summaryrefslogtreecommitdiff
path: root/src/Data
diff options
context:
space:
mode:
authorSean Whitton <spwhitton@spwhitton.name>2018-11-23 17:29:46 -0700
committerSean Whitton <spwhitton@spwhitton.name>2018-11-23 17:29:46 -0700
commit5a91775bd7c909bbaf3b3c8ee964136961a5146d (patch)
treecea89f085c1382f1567ff71a30b6f953708dc6a9 /src/Data
parent5d2b156377dce5bdca65b14639306eaed3ac3a92 (diff)
parenta19b25a865b2000bbd3acd909f5951a5407c1eec (diff)
Merge tag 'upstream/0.17'
Upstream version 0.17 # gpg: Signature made Fri 23 Nov 2018 05:29:18 PM MST # gpg: using RSA key 9B917007AE030E36E4FC248B695B7AE4BF066240 # gpg: issuer "spwhitton@spwhitton.name" # gpg: Good signature from "Sean Whitton <spwhitton@spwhitton.name>" [ultimate] # Primary key fingerprint: 8DC2 487E 51AB DD90 B5C4 753F 0F56 D055 3B6D 411B # Subkey fingerprint: 9B91 7007 AE03 0E36 E4FC 248B 695B 7AE4 BF06 6240
Diffstat (limited to 'src/Data')
-rw-r--r--src/Data/AVL.agda421
-rw-r--r--src/Data/AVL/Height.agda56
-rw-r--r--src/Data/AVL/Indexed.agda262
-rw-r--r--src/Data/AVL/IndexedMap.agda15
-rw-r--r--src/Data/AVL/Key.agda68
-rw-r--r--src/Data/AVL/Sets.agda5
-rw-r--r--src/Data/Bin.agda6
-rw-r--r--src/Data/Bin/Properties.agda10
-rw-r--r--src/Data/Bool/Base.agda1
-rw-r--r--src/Data/Bool/Properties.agda581
-rw-r--r--src/Data/Bool/Solver.agda28
-rw-r--r--src/Data/BoundedVec.agda4
-rw-r--r--src/Data/Char.agda54
-rw-r--r--src/Data/Char/Base.agda36
-rw-r--r--src/Data/Char/Unsafe.agda58
-rw-r--r--src/Data/Cofin.agda42
-rw-r--r--src/Data/Colist.agda501
-rw-r--r--src/Data/Colist/Infinite-merge.agda231
-rw-r--r--src/Data/Conat.agda64
-rw-r--r--src/Data/Container.agda236
-rw-r--r--src/Data/Container/Any.agda386
-rw-r--r--src/Data/Container/Combinator.agda142
-rw-r--r--src/Data/Container/Core.agda53
-rw-r--r--src/Data/Container/FreeMonad.agda30
-rw-r--r--src/Data/Container/Indexed.agda44
-rw-r--r--src/Data/Container/Indexed/Combinator.agda87
-rw-r--r--src/Data/Container/Indexed/FreeMonad.agda12
-rw-r--r--src/Data/Covec.agda155
-rw-r--r--src/Data/Digit.agda5
-rw-r--r--src/Data/Fin.agda235
-rw-r--r--src/Data/Fin/Base.agda239
-rw-r--r--src/Data/Fin/Dec.agda174
-rw-r--r--src/Data/Fin/Literals.agda18
-rw-r--r--src/Data/Fin/Permutation.agda183
-rw-r--r--src/Data/Fin/Permutation/Components.agda75
-rw-r--r--src/Data/Fin/Properties.agda694
-rw-r--r--src/Data/Fin/Subset.agda84
-rw-r--r--src/Data/Fin/Subset/Properties.agda538
-rw-r--r--src/Data/Fin/Substitution.agda28
-rw-r--r--src/Data/Fin/Substitution/Example.agda5
-rw-r--r--src/Data/Fin/Substitution/Lemmas.agda21
-rw-r--r--src/Data/Fin/Substitution/List.agda3
-rw-r--r--src/Data/Float.agda25
-rw-r--r--src/Data/Float/Unsafe.agda24
-rw-r--r--src/Data/Graph/Acyclic.agda167
-rw-r--r--src/Data/Integer.agda54
-rw-r--r--src/Data/Integer/Addition/Properties.agda4
-rw-r--r--src/Data/Integer/Base.agda55
-rw-r--r--src/Data/Integer/Divisibility.agda2
-rw-r--r--src/Data/Integer/Literals.agda24
-rw-r--r--src/Data/Integer/Properties.agda504
-rw-r--r--src/Data/Integer/Solver.agda22
-rw-r--r--src/Data/List.agda80
-rw-r--r--src/Data/List/All.agda55
-rw-r--r--src/Data/List/All/Properties.agda249
-rw-r--r--src/Data/List/Any.agda62
-rw-r--r--src/Data/List/Any/BagAndSetEquality.agda274
-rw-r--r--src/Data/List/Any/Membership.agda55
-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.agda463
-rw-r--r--src/Data/List/Base.agda320
-rw-r--r--src/Data/List/Categorical.agda252
-rw-r--r--src/Data/List/Countdown.agda19
-rw-r--r--src/Data/List/Literals.agda19
-rw-r--r--src/Data/List/Membership/DecPropositional.agda19
-rw-r--r--src/Data/List/Membership/DecSetoid.agda25
-rw-r--r--src/Data/List/Membership/Propositional.agda24
-rw-r--r--src/Data/List/Membership/Propositional/Properties.agda332
-rw-r--r--src/Data/List/Membership/Propositional/Properties/Core.agda74
-rw-r--r--src/Data/List/Membership/Setoid.agda59
-rw-r--r--src/Data/List/Membership/Setoid/Properties.agda239
-rw-r--r--src/Data/List/NonEmpty.agda111
-rw-r--r--src/Data/List/NonEmpty/Categorical.agda87
-rw-r--r--src/Data/List/NonEmpty/Properties.agda12
-rw-r--r--src/Data/List/Properties.agda1131
-rw-r--r--src/Data/List/Relation/BagAndSetEquality.agda325
-rw-r--r--src/Data/List/Relation/Equality/DecPropositional.agda22
-rw-r--r--src/Data/List/Relation/Equality/DecSetoid.agda34
-rw-r--r--src/Data/List/Relation/Equality/Propositional.agda28
-rw-r--r--src/Data/List/Relation/Equality/Setoid.agda58
-rw-r--r--src/Data/List/Relation/Lex/Core.agda106
-rw-r--r--src/Data/List/Relation/Lex/NonStrict.agda188
-rw-r--r--src/Data/List/Relation/Lex/Strict.agda227
-rw-r--r--src/Data/List/Relation/Permutation/Inductive.agda78
-rw-r--r--src/Data/List/Relation/Permutation/Inductive/Properties.agda274
-rw-r--r--src/Data/List/Relation/Pointwise.agda273
-rw-r--r--src/Data/List/Relation/Sublist/Propositional.agda48
-rw-r--r--src/Data/List/Relation/Sublist/Propositional/Properties.agda360
-rw-r--r--src/Data/List/Relation/Sublist/Propositional/Solver.agda144
-rw-r--r--src/Data/List/Relation/Subset/Propositional.agda16
-rw-r--r--src/Data/List/Relation/Subset/Propositional/Properties.agda208
-rw-r--r--src/Data/List/Relation/Subset/Setoid.agda28
-rw-r--r--src/Data/List/Relation/Subset/Setoid/Properties.agda78
-rw-r--r--src/Data/List/Reverse.agda37
-rw-r--r--src/Data/List/Solver.agda19
-rw-r--r--src/Data/List/Zipper.agda116
-rw-r--r--src/Data/List/Zipper/Properties.agda131
-rw-r--r--src/Data/M.agda25
-rw-r--r--src/Data/M/Indexed.agda42
-rw-r--r--src/Data/Maybe.agda38
-rw-r--r--src/Data/Maybe/Base.agda52
-rw-r--r--src/Data/Maybe/Categorical.agda81
-rw-r--r--src/Data/Nat.agda37
-rw-r--r--src/Data/Nat/Base.agda144
-rw-r--r--src/Data/Nat/Coprimality.agda53
-rw-r--r--src/Data/Nat/DivMod.agda203
-rw-r--r--src/Data/Nat/DivMod/Core.agda102
-rw-r--r--src/Data/Nat/DivMod/Unsafe.agda45
-rw-r--r--src/Data/Nat/Divisibility.agda207
-rw-r--r--src/Data/Nat/GCD.agda42
-rw-r--r--src/Data/Nat/GCD/Lemmas.agda76
-rw-r--r--src/Data/Nat/InfinitelyOften.agda41
-rw-r--r--src/Data/Nat/LCM.agda11
-rw-r--r--src/Data/Nat/Literals.agda17
-rw-r--r--src/Data/Nat/Primality.agda22
-rw-r--r--src/Data/Nat/Properties.agda763
-rw-r--r--src/Data/Nat/Solver.agda21
-rw-r--r--src/Data/Nat/Unsafe.agda13
-rw-r--r--src/Data/Plus.agda80
-rw-r--r--src/Data/Product.agda31
-rw-r--r--src/Data/Product/Categorical/Examples.agda60
-rw-r--r--src/Data/Product/Categorical/Left.agda79
-rw-r--r--src/Data/Product/Categorical/Left/Base.agda35
-rw-r--r--src/Data/Product/Categorical/Right.agda78
-rw-r--r--src/Data/Product/Categorical/Right/Base.agda35
-rw-r--r--src/Data/Product/N-ary.agda171
-rw-r--r--src/Data/Product/N-ary/Categorical.agda59
-rw-r--r--src/Data/Product/N-ary/Properties.agda89
-rw-r--r--src/Data/Product/Properties.agda19
-rw-r--r--src/Data/Product/Relation/Lex/NonStrict.agda213
-rw-r--r--src/Data/Product/Relation/Lex/Strict.agda302
-rw-r--r--src/Data/Product/Relation/Pointwise/Dependent.agda449
-rw-r--r--src/Data/Product/Relation/Pointwise/NonDependent.agda499
-rw-r--r--src/Data/Rational.agda18
-rw-r--r--src/Data/Rational/Literals.agda35
-rw-r--r--src/Data/Rational/Properties.agda8
-rw-r--r--src/Data/ReflexiveClosure.agda44
-rw-r--r--src/Data/Sign.agda7
-rw-r--r--src/Data/Sign/Properties.agda104
-rw-r--r--src/Data/Star.agda134
-rw-r--r--src/Data/Star/BoundedVec.agda4
-rw-r--r--src/Data/Star/Decoration.agda48
-rw-r--r--src/Data/Star/Environment.agda16
-rw-r--r--src/Data/Star/Fin.agda3
-rw-r--r--src/Data/Star/List.agda17
-rw-r--r--src/Data/Star/Nat.agda6
-rw-r--r--src/Data/Star/Pointer.agda84
-rw-r--r--src/Data/Star/Properties.agda93
-rw-r--r--src/Data/Star/Vec.agda22
-rw-r--r--src/Data/Stream.agda162
-rw-r--r--src/Data/String.agda68
-rw-r--r--src/Data/String/Base.agda57
-rw-r--r--src/Data/String/Literals.agda17
-rw-r--r--src/Data/String/Unsafe.agda65
-rw-r--r--src/Data/Sum.agda70
-rw-r--r--src/Data/Sum/Base.agda56
-rw-r--r--src/Data/Sum/Categorical/Examples.agda52
-rw-r--r--src/Data/Sum/Categorical/Left.agda76
-rw-r--r--src/Data/Sum/Categorical/Right.agda68
-rw-r--r--src/Data/Sum/Properties.agda22
-rw-r--r--src/Data/Sum/Relation/Core.agda132
-rw-r--r--src/Data/Sum/Relation/LeftOrder.agda220
-rw-r--r--src/Data/Sum/Relation/Pointwise.agda389
-rw-r--r--src/Data/Table.agda36
-rw-r--r--src/Data/Table/Base.agda93
-rw-r--r--src/Data/Table/Properties.agda117
-rw-r--r--src/Data/Table/Relation/Equality.agda33
-rw-r--r--src/Data/These.agda77
-rw-r--r--src/Data/These/Categorical/Left.agda54
-rw-r--r--src/Data/These/Categorical/Left/Base.agda59
-rw-r--r--src/Data/These/Categorical/Right.agda54
-rw-r--r--src/Data/These/Categorical/Right/Base.agda59
-rw-r--r--src/Data/Vec.agda237
-rw-r--r--src/Data/Vec/All.agda81
-rw-r--r--src/Data/Vec/All/Properties.agda287
-rw-r--r--src/Data/Vec/Any.agda79
-rw-r--r--src/Data/Vec/Categorical.agda81
-rw-r--r--src/Data/Vec/Equality.agda105
-rw-r--r--src/Data/Vec/Membership/Propositional.agda24
-rw-r--r--src/Data/Vec/Membership/Propositional/Properties.agda103
-rw-r--r--src/Data/Vec/Properties.agda1013
-rw-r--r--src/Data/Vec/Relation/Equality/DecPropositional.agda22
-rw-r--r--src/Data/Vec/Relation/Equality/DecSetoid.agda37
-rw-r--r--src/Data/Vec/Relation/Equality/Propositional.agda36
-rw-r--r--src/Data/Vec/Relation/Equality/Setoid.agda87
-rw-r--r--src/Data/Vec/Relation/Pointwise/Extensional.agda227
-rw-r--r--src/Data/Vec/Relation/Pointwise/Inductive.agda258
-rw-r--r--src/Data/W.agda60
-rw-r--r--src/Data/Word.agda (renamed from src/Data/Char/Core.agda)13
-rw-r--r--src/Data/Word/Unsafe.agda22
192 files changed, 16189 insertions, 7468 deletions
diff --git a/src/Data/AVL.agda b/src/Data/AVL.agda
index f8bb93b..e31de69 100644
--- a/src/Data/AVL.agda
+++ b/src/Data/AVL.agda
@@ -10,12 +10,10 @@
-- described by Conor McBride in his talk "Pivotal pragmatism".
open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Relation.Binary.PropositionalEquality as P using (_≡_ ; refl)
module Data.AVL
- {k v ℓ}
- {Key : Set k} (Value : Key → Set v)
- {_<_ : Rel Key ℓ}
+ {k r} {Key : Set k} {_<_ : Rel Key r}
(isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_)
where
@@ -31,392 +29,85 @@ open import Function
open import Level using (_⊔_; Lift; lift)
open IsStrictTotalOrder isStrictTotalOrder
+import Data.AVL.Indexed Key isStrictTotalOrder as Indexed
+open Indexed using (K&_ ; ⊥⁺ ; ⊤⁺)
------------------------------------------------------------------------
--- Extended keys
-
-module Extended-key where
-
- -- The key type extended with a new minimum and maximum.
-
- data Key⁺ : Set k where
- ⊥⁺ ⊤⁺ : Key⁺
- [_] : (k : Key) → Key⁺
-
- -- An extended strict ordering relation.
-
- infix 4 _<⁺_
-
- _<⁺_ : Key⁺ → Key⁺ → Set ℓ
- ⊥⁺ <⁺ [ _ ] = Lift ⊤
- ⊥⁺ <⁺ ⊤⁺ = Lift ⊤
- [ x ] <⁺ [ y ] = x < y
- [ _ ] <⁺ ⊤⁺ = Lift ⊤
- _ <⁺ _ = Lift ⊥
-
- -- A pair of ordering constraints.
-
- infix 4 _<_<_
-
- _<_<_ : Key⁺ → Key → Key⁺ → Set ℓ
- l < x < u = l <⁺ [ x ] × [ x ] <⁺ u
-
- -- _<⁺_ is transitive.
-
- trans⁺ : ∀ l {m u} → l <⁺ m → m <⁺ u → l <⁺ u
-
- trans⁺ [ l ] {m = [ m ]} {u = [ u ]} l<m m<u = trans l<m m<u
-
- trans⁺ ⊥⁺ {u = [ _ ]} _ _ = _
- trans⁺ ⊥⁺ {u = ⊤⁺} _ _ = _
- trans⁺ [ _ ] {u = ⊤⁺} _ _ = _
-
- trans⁺ _ {m = ⊥⁺} {u = ⊥⁺} _ (lift ())
- trans⁺ _ {m = [ _ ]} {u = ⊥⁺} _ (lift ())
- trans⁺ _ {m = ⊤⁺} {u = ⊥⁺} _ (lift ())
- trans⁺ [ _ ] {m = ⊥⁺} {u = [ _ ]} (lift ()) _
- trans⁺ [ _ ] {m = ⊤⁺} {u = [ _ ]} _ (lift ())
- trans⁺ ⊤⁺ {m = ⊥⁺} (lift ()) _
- trans⁺ ⊤⁺ {m = [ _ ]} (lift ()) _
- trans⁺ ⊤⁺ {m = ⊤⁺} (lift ()) _
-
-------------------------------------------------------------------------
--- Types and functions which are used to keep track of height
--- invariants
-
-module Height-invariants where
-
- -- Bits. (I would use Fin 2 instead if Agda had "defined patterns",
- -- so that I could pattern match on 1# instead of suc zero; the text
- -- "suc zero" takes up a lot more space.)
-
- data ℕ₂ : Set where
- 0# : ℕ₂
- 1# : ℕ₂
-
- -- Addition.
-
- infixl 6 _⊕_
-
- _⊕_ : ℕ₂ → ℕ → ℕ
- 0# ⊕ n = n
- 1# ⊕ n = 1 + n
-
- -- pred[ i ⊕ n ] = pred (i ⊕ n).
-
- pred[_⊕_] : ℕ₂ → ℕ → ℕ
- pred[ i ⊕ zero ] = 0
- pred[ i ⊕ suc n ] = i ⊕ n
-
- infix 4 _∼_⊔_
-
- -- If i ∼ j ⊔ m, then the difference between i and j is at most 1,
- -- and the maximum of i and j is m. _∼_⊔_ is used to record the
- -- balance factor of the AVL trees, and also to ensure that the
- -- absolute value of the balance factor is never more than 1.
-
- data _∼_⊔_ : ℕ → ℕ → ℕ → Set where
- ∼+ : ∀ {n} → n ∼ 1 + n ⊔ 1 + n
- ∼0 : ∀ {n} → n ∼ n ⊔ n
- ∼- : ∀ {n} → 1 + n ∼ n ⊔ 1 + n
-
- -- Some lemmas.
-
- max∼ : ∀ {i j m} → i ∼ j ⊔ m → m ∼ i ⊔ m
- max∼ ∼+ = ∼-
- max∼ ∼0 = ∼0
- max∼ ∼- = ∼0
-
- ∼max : ∀ {i j m} → i ∼ j ⊔ m → j ∼ m ⊔ m
- ∼max ∼+ = ∼0
- ∼max ∼0 = ∼0
- ∼max ∼- = ∼+
+-- Types and functions with hidden indices
-------------------------------------------------------------------------
--- AVL trees
+data Tree {v} (V : Key → Set v) : Set (k ⊔ v ⊔ r) where
+ tree : ∀ {h} → Indexed.Tree V ⊥⁺ ⊤⁺ h → Tree V
--- Key/value pairs.
-
-KV : Set (k ⊔ v)
-KV = Σ Key Value
-
-module Indexed where
-
- open Extended-key
- open Height-invariants
-
- -- The trees have three parameters/indices: a lower bound on the
- -- keys, an upper bound, and a height.
- --
- -- (The bal argument is the balance factor.)
-
- data Tree (l u : Key⁺) : ℕ → Set (k ⊔ v ⊔ ℓ) where
- leaf : (l<u : l <⁺ u) → Tree l u 0
- node : ∀ {hˡ hʳ h}
- (k : KV)
- (lk : Tree l [ proj₁ k ] hˡ)
- (ku : Tree [ proj₁ k ] u hʳ)
- (bal : hˡ ∼ hʳ ⊔ h) →
- Tree l u (suc h)
-
- -- Cast operations. Logarithmic in the size of the tree, if we don't
- -- count the time needed to construct the new proofs in the leaf
- -- cases. (The same kind of caveat applies to other operations
- -- below.)
- --
- -- Perhaps it would be worthwhile changing the data structure so
- -- that the casts could be implemented in constant time (excluding
- -- proof manipulation). However, note that this would not change the
- -- worst-case time complexity of the operations below (up to Θ).
-
- castˡ : ∀ {l m u h} → l <⁺ m → Tree m u h → Tree l u h
- castˡ {l} l<m (leaf m<u) = leaf (trans⁺ l l<m m<u)
- castˡ l<m (node k mk ku bal) = node k (castˡ l<m mk) ku bal
-
- castʳ : ∀ {l m u h} → Tree l m h → m <⁺ u → Tree l u h
- castʳ {l} (leaf l<m) m<u = leaf (trans⁺ l l<m m<u)
- castʳ (node k lk km bal) m<u = node k lk (castʳ km m<u) bal
-
- -- Various constant-time functions which construct trees out of
- -- smaller pieces, sometimes using rotation.
-
- joinˡ⁺ : ∀ {l u hˡ hʳ h} →
- (k : KV) →
- (∃ λ i → Tree l [ proj₁ k ] (i ⊕ hˡ)) →
- Tree [ proj₁ k ] u hʳ →
- (bal : hˡ ∼ hʳ ⊔ h) →
- ∃ λ i → Tree l u (i ⊕ (1 + h))
- joinˡ⁺ k₆ (1# , node k₂ t₁
- (node k₄ t₃ t₅ bal)
- ∼+) t₇ ∼- = (0# , node k₄
- (node k₂ t₁ t₃ (max∼ bal))
- (node k₆ t₅ t₇ (∼max bal))
- ∼0)
- joinˡ⁺ k₄ (1# , node k₂ t₁ t₃ ∼-) t₅ ∼- = (0# , node k₂ t₁ (node k₄ t₃ t₅ ∼0) ∼0)
- joinˡ⁺ k₄ (1# , node k₂ t₁ t₃ ∼0) t₅ ∼- = (1# , node k₂ t₁ (node k₄ t₃ t₅ ∼-) ∼+)
- joinˡ⁺ k₂ (1# , t₁) t₃ ∼0 = (1# , node k₂ t₁ t₃ ∼-)
- joinˡ⁺ k₂ (1# , t₁) t₃ ∼+ = (0# , node k₂ t₁ t₃ ∼0)
- joinˡ⁺ k₂ (0# , t₁) t₃ bal = (0# , node k₂ t₁ t₃ bal)
-
- joinʳ⁺ : ∀ {l u hˡ hʳ h} →
- (k : KV) →
- Tree l [ proj₁ k ] hˡ →
- (∃ λ i → Tree [ proj₁ k ] u (i ⊕ hʳ)) →
- (bal : hˡ ∼ hʳ ⊔ h) →
- ∃ λ i → Tree l u (i ⊕ (1 + h))
- joinʳ⁺ k₂ t₁ (1# , node k₆
- (node k₄ t₃ t₅ bal)
- t₇ ∼-) ∼+ = (0# , node k₄
- (node k₂ t₁ t₃ (max∼ bal))
- (node k₆ t₅ t₇ (∼max bal))
- ∼0)
- joinʳ⁺ k₂ t₁ (1# , node k₄ t₃ t₅ ∼+) ∼+ = (0# , node k₄ (node k₂ t₁ t₃ ∼0) t₅ ∼0)
- joinʳ⁺ k₂ t₁ (1# , node k₄ t₃ t₅ ∼0) ∼+ = (1# , node k₄ (node k₂ t₁ t₃ ∼+) t₅ ∼-)
- joinʳ⁺ k₂ t₁ (1# , t₃) ∼0 = (1# , node k₂ t₁ t₃ ∼+)
- joinʳ⁺ k₂ t₁ (1# , t₃) ∼- = (0# , node k₂ t₁ t₃ ∼0)
- joinʳ⁺ k₂ t₁ (0# , t₃) bal = (0# , node k₂ t₁ t₃ bal)
-
- joinˡ⁻ : ∀ {l u} hˡ {hʳ h} →
- (k : KV) →
- (∃ λ i → Tree l [ proj₁ k ] pred[ i ⊕ hˡ ]) →
- Tree [ proj₁ k ] u hʳ →
- (bal : hˡ ∼ hʳ ⊔ h) →
- ∃ λ i → Tree l u (i ⊕ h)
- joinˡ⁻ zero k₂ (0# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
- joinˡ⁻ zero k₂ (1# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
- joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼+ = joinʳ⁺ k₂ t₁ (1# , t₃) ∼+
- joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼0 = (1# , node k₂ t₁ t₃ ∼+)
- joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼- = (0# , node k₂ t₁ t₃ ∼0)
- joinˡ⁻ (suc _) k₂ (1# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
-
- joinʳ⁻ : ∀ {l u hˡ} hʳ {h} →
- (k : KV) →
- Tree l [ proj₁ k ] hˡ →
- (∃ λ i → Tree [ proj₁ k ] u pred[ i ⊕ hʳ ]) →
- (bal : hˡ ∼ hʳ ⊔ h) →
- ∃ λ i → Tree l u (i ⊕ h)
- joinʳ⁻ zero k₂ t₁ (0# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
- joinʳ⁻ zero k₂ t₁ (1# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
- joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼- = joinˡ⁺ k₂ (1# , t₁) t₃ ∼-
- joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼0 = (1# , node k₂ t₁ t₃ ∼-)
- joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼+ = (0# , node k₂ t₁ t₃ ∼0)
- joinʳ⁻ (suc _) k₂ t₁ (1# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
-
- -- Extracts the smallest element from the tree, plus the rest.
- -- Logarithmic in the size of the tree.
-
- headTail : ∀ {l u h} → Tree l u (1 + h) →
- ∃ λ (k : KV) → l <⁺ [ proj₁ k ] ×
- ∃ λ i → Tree [ proj₁ k ] u (i ⊕ h)
- headTail (node k₁ (leaf l<k₁) t₂ ∼+) = (k₁ , l<k₁ , 0# , t₂)
- headTail (node k₁ (leaf l<k₁) t₂ ∼0) = (k₁ , l<k₁ , 0# , t₂)
- headTail (node {hˡ = suc _} k₃ t₁₂ t₄ bal) with headTail t₁₂
- ... | (k₁ , l<k₁ , t₂) = (k₁ , l<k₁ , joinˡ⁻ _ k₃ t₂ t₄ bal)
-
- -- Extracts the largest element from the tree, plus the rest.
- -- Logarithmic in the size of the tree.
-
- initLast : ∀ {l u h} → Tree l u (1 + h) →
- ∃ λ (k : KV) → [ proj₁ k ] <⁺ u ×
- ∃ λ i → Tree l [ proj₁ k ] (i ⊕ h)
- initLast (node k₂ t₁ (leaf k₂<u) ∼-) = (k₂ , k₂<u , (0# , t₁))
- initLast (node k₂ t₁ (leaf k₂<u) ∼0) = (k₂ , k₂<u , (0# , t₁))
- initLast (node {hʳ = suc _} k₂ t₁ t₃₄ bal) with initLast t₃₄
- ... | (k₄ , k₄<u , t₃) = (k₄ , k₄<u , joinʳ⁻ _ k₂ t₁ t₃ bal)
-
- -- Another joining function. Logarithmic in the size of either of
- -- the input trees (which need to have almost equal heights).
-
- join : ∀ {l m u hˡ hʳ h} →
- Tree l m hˡ → Tree m u hʳ → (bal : hˡ ∼ hʳ ⊔ h) →
- ∃ λ i → Tree l u (i ⊕ h)
- join t₁ (leaf m<u) ∼0 = (0# , castʳ t₁ m<u)
- join t₁ (leaf m<u) ∼- = (0# , castʳ t₁ m<u)
- join {hʳ = suc _} t₁ t₂₃ bal with headTail t₂₃
- ... | (k₂ , m<k₂ , t₃) = joinʳ⁻ _ k₂ (castʳ t₁ m<k₂) t₃ bal
-
- -- An empty tree.
-
- empty : ∀ {l u} → l <⁺ u → Tree l u 0
- empty = leaf
-
- -- A singleton tree.
-
- singleton : ∀ {l u} (k : Key) → Value k → l < k < u → Tree l u 1
- singleton k v (l<k , k<u) = node (k , v) (leaf l<k) (leaf k<u) ∼0
-
- -- Inserts a key into the tree, using a function to combine any
- -- existing value with the new value. Logarithmic in the size of the
- -- tree (assuming constant-time comparisons and a constant-time
- -- combining function).
-
- insertWith : ∀ {l u h} → (k : Key) → Value k →
- (Value k → Value k → Value k) → -- New → old → result.
- Tree l u h → l < k < u →
- ∃ λ i → Tree l u (i ⊕ h)
- insertWith k v f (leaf l<u) l<k<u = (1# , singleton k v l<k<u)
- insertWith k v f (node (k′ , v′) lp pu bal) (l<k , k<u) with compare k k′
- ... | tri< k<k′ _ _ = joinˡ⁺ (k′ , v′) (insertWith k v f lp (l<k , k<k′)) pu bal
- ... | tri> _ _ k′<k = joinʳ⁺ (k′ , v′) lp (insertWith k v f pu (k′<k , k<u)) bal
- ... | tri≈ _ k≡k′ _ rewrite P.sym k≡k′ = (0# , node (k , f v v′) lp pu bal)
-
- -- Inserts a key into the tree. If the key already exists, then it
- -- is replaced. Logarithmic in the size of the tree (assuming
- -- constant-time comparisons).
-
- insert : ∀ {l u h} → (k : Key) → Value k → Tree l u h → l < k < u →
- ∃ λ i → Tree l u (i ⊕ h)
- insert k v = insertWith k v const
-
- -- Deletes the key/value pair containing the given key, if any.
- -- Logarithmic in the size of the tree (assuming constant-time
- -- comparisons).
-
- delete : ∀ {l u h} → Key → Tree l u h →
- ∃ λ i → Tree l u pred[ i ⊕ h ]
- delete k (leaf l<u) = (0# , leaf l<u)
- delete k (node p lp pu bal) with compare k (proj₁ p)
- ... | tri< _ _ _ = joinˡ⁻ _ p (delete k lp) pu bal
- ... | tri> _ _ _ = joinʳ⁻ _ p lp (delete k pu) bal
- ... | tri≈ _ _ _ = join lp pu bal
-
- -- Looks up a key. Logarithmic in the size of the tree (assuming
- -- constant-time comparisons).
-
- lookup : ∀ {l u h} → (k : Key) → Tree l u h → Maybe (Value k)
- lookup k (leaf _) = nothing
- lookup k (node (k′ , v) lk′ k′u _) with compare k k′
- ... | tri< _ _ _ = lookup k lk′
- ... | tri> _ _ _ = lookup k k′u
- ... | tri≈ _ eq _ rewrite eq = just v
-
- -- Maps a function over all values in the tree.
-
- map : (∀ {k} → Value k → Value k) →
- ∀ {l u h} → Tree l u h → Tree l u h
- map f (leaf l<u) = leaf l<u
- map f (node (k , v) l r bal) = node (k , f v) (map f l) (map f r) bal
-
- -- Converts the tree to an ordered list. Linear in the size of the
- -- tree.
-
- open DiffList
-
- toDiffList : ∀ {l u h} → Tree l u h → DiffList KV
- toDiffList (leaf _) = []
- toDiffList (node k l r _) = toDiffList l ++ k ∷ toDiffList r
+module _ {v} {V : Key → Set v} where
-------------------------------------------------------------------------
--- Types and functions with hidden indices
+ empty : Tree V
+ empty = tree (Indexed.empty _)
-data Tree : Set (k ⊔ v ⊔ ℓ) where
- tree : let open Extended-key in
- ∀ {h} → Indexed.Tree ⊥⁺ ⊤⁺ h → Tree
+ singleton : (k : Key) → V k → Tree V
+ singleton k v = tree (Indexed.singleton k v _)
-empty : Tree
-empty = tree (Indexed.empty _)
+ insert : (k : Key) → V k → Tree V → Tree V
+ insert k v (tree t) = tree $ proj₂ $ Indexed.insert k v t _
-singleton : (k : Key) → Value k → Tree
-singleton k v = tree (Indexed.singleton k v _)
+ insertWith : (k : Key) → V k → (V k → V k → V k) →
+ Tree V → Tree V
+ insertWith k v f (tree t) = tree $ proj₂ $ Indexed.insertWith k v f t _
-insert : (k : Key) → Value k → Tree → Tree
-insert k v (tree t) = tree $ proj₂ $ Indexed.insert k v t _
+ delete : Key → Tree V → Tree V
+ delete k (tree t) = tree $ proj₂ $ Indexed.delete k t
-insertWith : (k : Key) → Value k → (Value k → Value k → Value k) →
- Tree → Tree
-insertWith k v f (tree t) = tree $ proj₂ $ Indexed.insertWith k v f t _
+ lookup : (k : Key) → Tree V → Maybe (V k)
+ lookup k (tree t) = Indexed.lookup k t
-delete : Key → Tree → Tree
-delete k (tree t) = tree $ proj₂ $ Indexed.delete k t
+module _ {v w} {V : Key → Set v} {W : Key → Set w} where
-lookup : (k : Key) → Tree → Maybe (Value k)
-lookup k (tree t) = Indexed.lookup k t
+ map : ({k : Key} → V k → W k) → Tree V → Tree W
+ map f (tree t) = tree $ Indexed.map f t
-map : ({k : Key} → Value k → Value k) → Tree → Tree
-map f (tree t) = tree $ Indexed.map f t
+module _ {v} {V : Key → Set v} where
-infix 4 _∈?_
+ infix 4 _∈?_
-_∈?_ : Key → Tree → Bool
-k ∈? t = is-just (lookup k t)
+ _∈?_ : Key → Tree V → Bool
+ k ∈? t = is-just (lookup k t)
-headTail : Tree → Maybe (KV × Tree)
-headTail (tree (Indexed.leaf _)) = nothing
-headTail (tree {h = suc _} t) with Indexed.headTail t
-... | (k , _ , _ , t′) = just (k , tree (Indexed.castˡ _ t′))
+ headTail : Tree V → Maybe ((K& V) × Tree V)
+ headTail (tree (Indexed.leaf _)) = nothing
+ headTail (tree {h = suc _} t) with Indexed.headTail t
+ ... | (k , _ , _ , t′) = just (k , tree (Indexed.castˡ _ t′))
-initLast : Tree → Maybe (Tree × KV)
-initLast (tree (Indexed.leaf _)) = nothing
-initLast (tree {h = suc _} t) with Indexed.initLast t
-... | (k , _ , _ , t′) = just (tree (Indexed.castʳ t′ _) , k)
+ initLast : Tree V → Maybe (Tree V × (K& V))
+ initLast (tree (Indexed.leaf _)) = nothing
+ initLast (tree {h = suc _} t) with Indexed.initLast t
+ ... | (k , _ , _ , t′) = just (tree (Indexed.castʳ t′ _) , k)
--- The input does not need to be ordered.
+ -- The input does not need to be ordered.
-fromList : List KV → Tree
-fromList = List.foldr (uncurry insert) empty
+ fromList : List (K& V) → Tree V
+ fromList = List.foldr (uncurry insert) empty
--- Returns an ordered list.
+ -- Returns an ordered list.
-toList : Tree → List KV
-toList (tree t) = DiffList.toList (Indexed.toDiffList t)
+ toList : Tree V → List (K& V)
+ toList (tree t) = DiffList.toList (Indexed.toDiffList t)
--- Naive implementations of union.
+ -- Naive implementations of union.
-unionWith : (∀ {k} → Value k → Value k → Value k) →
- -- Left → right → result.
- Tree → Tree → Tree
-unionWith f t₁ t₂ =
- List.foldr (λ { (k , v) → insertWith k v f }) t₂ (toList t₁)
+ unionWith : (∀ {k} → V k → V k → V k) →
+ -- Left → right → result.
+ Tree V → Tree V → Tree V
+ unionWith f t₁ t₂ =
+ List.foldr (λ { (k , v) → insertWith k v f }) t₂ (toList t₁)
--- Left-biased.
+ -- Left-biased.
-union : Tree → Tree → Tree
-union = unionWith const
+ union : Tree V → Tree V → Tree V
+ union = unionWith const
-unionsWith : (∀ {k} → Value k → Value k → Value k) → List Tree → Tree
-unionsWith f ts = List.foldr (unionWith f) empty ts
+ unionsWith : (∀ {k} → V k → V k → V k) → List (Tree V) → Tree V
+ unionsWith f ts = List.foldr (unionWith f) empty ts
--- Left-biased.
+ -- Left-biased.
-unions : List Tree → Tree
-unions = unionsWith const
+ unions : List (Tree V) → Tree V
+ unions = unionsWith const
diff --git a/src/Data/AVL/Height.agda b/src/Data/AVL/Height.agda
new file mode 100644
index 0000000..6fe75ef
--- /dev/null
+++ b/src/Data/AVL/Height.agda
@@ -0,0 +1,56 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Types and functions which are used to keep track of height
+-- invariants in AVL Trees
+------------------------------------------------------------------------
+
+module Data.AVL.Height where
+
+open import Data.Nat.Base
+import Data.Fin as Fin
+
+ℕ₂ = Fin.Fin 2
+pattern 0# = Fin.zero
+pattern 1# = Fin.suc Fin.zero
+pattern ## = Fin.suc (Fin.suc ())
+
+-- Addition.
+
+infixl 6 _⊕_
+
+_⊕_ : ℕ₂ → ℕ → ℕ
+0# ⊕ n = n
+1# ⊕ n = 1 + n
+## ⊕ n
+
+-- pred[ i ⊕ n ] = pred (i ⊕ n).
+
+pred[_⊕_] : ℕ₂ → ℕ → ℕ
+pred[ i ⊕ zero ] = 0
+pred[ i ⊕ suc n ] = i ⊕ n
+
+infix 4 _∼_⊔_
+
+-- If i ∼ j ⊔ m, then the difference between i and j is at most 1,
+-- and the maximum of i and j is m. _∼_⊔_ is used to record the
+-- balance factor of the AVL trees, and also to ensure that the
+-- absolute value of the balance factor is never more than 1.
+
+data _∼_⊔_ : ℕ → ℕ → ℕ → Set where
+ ∼+ : ∀ {n} → n ∼ 1 + n ⊔ 1 + n
+ ∼0 : ∀ {n} → n ∼ n ⊔ n
+ ∼- : ∀ {n} → 1 + n ∼ n ⊔ 1 + n
+
+-- Some lemmas.
+
+max∼ : ∀ {i j m} → i ∼ j ⊔ m → m ∼ i ⊔ m
+max∼ ∼+ = ∼-
+max∼ ∼0 = ∼0
+max∼ ∼- = ∼0
+
+∼max : ∀ {i j m} → i ∼ j ⊔ m → j ∼ m ⊔ m
+∼max ∼+ = ∼0
+∼max ∼0 = ∼0
+∼max ∼- = ∼+
+
diff --git a/src/Data/AVL/Indexed.agda b/src/Data/AVL/Indexed.agda
new file mode 100644
index 0000000..6e4d579
--- /dev/null
+++ b/src/Data/AVL/Indexed.agda
@@ -0,0 +1,262 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Indexed AVL trees
+------------------------------------------------------------------------
+
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_ ; refl)
+
+module Data.AVL.Indexed
+ {k r} (Key : Set k) {_<_ : Rel Key r}
+ (isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_) where
+
+open import Level using (_⊔_) -- ; Lift; lift)
+open import Data.Nat.Base hiding (_<_; _⊔_; compare)
+open import Data.Product hiding (map)
+open import Data.Maybe hiding (map)
+import Data.DifferenceList as DiffList
+open import Function
+
+open IsStrictTotalOrder isStrictTotalOrder
+open import Data.AVL.Key Key isStrictTotalOrder public
+open import Data.AVL.Height public
+
+K&_ : ∀ {v} (Value : Key → Set v) → Set (k ⊔ v)
+K& Value = Σ Key Value
+
+-- The trees have three parameters/indices: a lower bound on the
+-- keys, an upper bound, and a height.
+--
+-- (The bal argument is the balance factor.)
+
+data Tree {v} (V : Key → Set v) (l u : Key⁺) : ℕ → Set (k ⊔ v ⊔ r) where
+ leaf : (l<u : l <⁺ u) → Tree V l u 0
+ node : ∀ {hˡ hʳ h}
+ (k : K& V)
+ (lk : Tree V l [ proj₁ k ] hˡ)
+ (ku : Tree V [ proj₁ k ] u hʳ)
+ (bal : hˡ ∼ hʳ ⊔ h) →
+ Tree V l u (suc h)
+
+module _ {v} {V : Key → Set v} where
+
+ leaf-injective : ∀ {l u} {p q : l <⁺ u} → (Tree V l u 0 ∋ leaf p) ≡ leaf q → p ≡ q
+ leaf-injective refl = refl
+
+ node-injective-key : ∀ {hˡ hʳ h l u k₁ k₂}
+ {lk₁ : Tree V l [ proj₁ k₁ ] hˡ} {lk₂ : Tree V l [ proj₁ k₂ ] hˡ}
+ {ku₁ : Tree V [ proj₁ k₁ ] u hʳ} {ku₂ : Tree V [ proj₁ k₂ ] u hʳ}
+ {bal₁ bal₂ : hˡ ∼ hʳ ⊔ h} → node k₁ lk₁ ku₁ bal₁ ≡ node k₂ lk₂ ku₂ bal₂ → k₁ ≡ k₂
+ node-injective-key refl = refl
+
+ node-injectiveˡ : ∀ {hˡ hʳ h l u k}
+ {lk₁ : Tree V l [ proj₁ k ] hˡ} {lk₂ : Tree V l [ proj₁ k ] hˡ}
+ {ku₁ : Tree V [ proj₁ k ] u hʳ} {ku₂ : Tree V [ proj₁ k ] u hʳ}
+ {bal₁ bal₂ : hˡ ∼ hʳ ⊔ h} → node k lk₁ ku₁ bal₁ ≡ node k lk₂ ku₂ bal₂ → lk₁ ≡ lk₂
+ node-injectiveˡ refl = refl
+
+ node-injectiveʳ : ∀ {hˡ hʳ h l u k}
+ {lk₁ : Tree V l [ proj₁ k ] hˡ} {lk₂ : Tree V l [ proj₁ k ] hˡ}
+ {ku₁ : Tree V [ proj₁ k ] u hʳ} {ku₂ : Tree V [ proj₁ k ] u hʳ}
+ {bal₁ bal₂ : hˡ ∼ hʳ ⊔ h} → node k lk₁ ku₁ bal₁ ≡ node k lk₂ ku₂ bal₂ → ku₁ ≡ ku₂
+ node-injectiveʳ refl = refl
+
+ node-injective-bal : ∀ {hˡ hʳ h l u k}
+ {lk₁ : Tree V l [ proj₁ k ] hˡ} {lk₂ : Tree V l [ proj₁ k ] hˡ}
+ {ku₁ : Tree V [ proj₁ k ] u hʳ} {ku₂ : Tree V [ proj₁ k ] u hʳ}
+ {bal₁ bal₂ : hˡ ∼ hʳ ⊔ h} → node k lk₁ ku₁ bal₁ ≡ node k lk₂ ku₂ bal₂ → bal₁ ≡ bal₂
+ node-injective-bal refl = refl
+
+ -- Cast operations. Logarithmic in the size of the tree, if we don't
+ -- count the time needed to construct the new proofs in the leaf
+ -- cases. (The same kind of caveat applies to other operations
+ -- below.)
+ --
+ -- Perhaps it would be worthwhile changing the data structure so
+ -- that the casts could be implemented in constant time (excluding
+ -- proof manipulation). However, note that this would not change the
+ -- worst-case time complexity of the operations below (up to Θ).
+
+ castˡ : ∀ {l m u h} → l <⁺ m → Tree V m u h → Tree V l u h
+ castˡ {l} l<m (leaf m<u) = leaf (trans⁺ l l<m m<u)
+ castˡ l<m (node k mk ku bal) = node k (castˡ l<m mk) ku bal
+
+ castʳ : ∀ {l m u h} → Tree V l m h → m <⁺ u → Tree V l u h
+ castʳ {l} (leaf l<m) m<u = leaf (trans⁺ l l<m m<u)
+ castʳ (node k lk km bal) m<u = node k lk (castʳ km m<u) bal
+
+ -- Various constant-time functions which construct trees out of
+ -- smaller pieces, sometimes using rotation.
+
+ joinˡ⁺ : ∀ {l u hˡ hʳ h} →
+ (k : K& V) →
+ (∃ λ i → Tree V l [ proj₁ k ] (i ⊕ hˡ)) →
+ Tree V [ proj₁ k ] u hʳ →
+ (bal : hˡ ∼ hʳ ⊔ h) →
+ ∃ λ i → Tree V l u (i ⊕ (1 + h))
+ joinˡ⁺ k₆ (1# , node k₂ t₁
+ (node k₄ t₃ t₅ bal)
+ ∼+) t₇ ∼- = (0# , node k₄
+ (node k₂ t₁ t₃ (max∼ bal))
+ (node k₆ t₅ t₇ (∼max bal))
+ ∼0)
+ joinˡ⁺ k₄ (1# , node k₂ t₁ t₃ ∼-) t₅ ∼- = (0# , node k₂ t₁ (node k₄ t₃ t₅ ∼0) ∼0)
+ joinˡ⁺ k₄ (1# , node k₂ t₁ t₃ ∼0) t₅ ∼- = (1# , node k₂ t₁ (node k₄ t₃ t₅ ∼-) ∼+)
+ joinˡ⁺ k₂ (1# , t₁) t₃ ∼0 = (1# , node k₂ t₁ t₃ ∼-)
+ joinˡ⁺ k₂ (1# , t₁) t₃ ∼+ = (0# , node k₂ t₁ t₃ ∼0)
+ joinˡ⁺ k₂ (0# , t₁) t₃ bal = (0# , node k₂ t₁ t₃ bal)
+ joinˡ⁺ k₂ (## , t₁) t₃ bal
+
+ joinʳ⁺ : ∀ {l u hˡ hʳ h} →
+ (k : K& V) →
+ Tree V l [ proj₁ k ] hˡ →
+ (∃ λ i → Tree V [ proj₁ k ] u (i ⊕ hʳ)) →
+ (bal : hˡ ∼ hʳ ⊔ h) →
+ ∃ λ i → Tree V l u (i ⊕ (1 + h))
+ joinʳ⁺ k₂ t₁ (1# , node k₆
+ (node k₄ t₃ t₅ bal)
+ t₇ ∼-) ∼+ = (0# , node k₄
+ (node k₂ t₁ t₃ (max∼ bal))
+ (node k₆ t₅ t₇ (∼max bal))
+ ∼0)
+ joinʳ⁺ k₂ t₁ (1# , node k₄ t₃ t₅ ∼+) ∼+ = (0# , node k₄ (node k₂ t₁ t₃ ∼0) t₅ ∼0)
+ joinʳ⁺ k₂ t₁ (1# , node k₄ t₃ t₅ ∼0) ∼+ = (1# , node k₄ (node k₂ t₁ t₃ ∼+) t₅ ∼-)
+ joinʳ⁺ k₂ t₁ (1# , t₃) ∼0 = (1# , node k₂ t₁ t₃ ∼+)
+ joinʳ⁺ k₂ t₁ (1# , t₃) ∼- = (0# , node k₂ t₁ t₃ ∼0)
+ joinʳ⁺ k₂ t₁ (0# , t₃) bal = (0# , node k₂ t₁ t₃ bal)
+ joinʳ⁺ k₂ t₁ (## , t₃) bal
+
+ joinˡ⁻ : ∀ {l u} hˡ {hʳ h} →
+ (k : K& V) →
+ (∃ λ i → Tree V l [ proj₁ k ] pred[ i ⊕ hˡ ]) →
+ Tree V [ proj₁ k ] u hʳ →
+ (bal : hˡ ∼ hʳ ⊔ h) →
+ ∃ λ i → Tree V l u (i ⊕ h)
+ joinˡ⁻ zero k₂ (0# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
+ joinˡ⁻ zero k₂ (1# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
+ joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼+ = joinʳ⁺ k₂ t₁ (1# , t₃) ∼+
+ joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼0 = (1# , node k₂ t₁ t₃ ∼+)
+ joinˡ⁻ (suc _) k₂ (0# , t₁) t₃ ∼- = (0# , node k₂ t₁ t₃ ∼0)
+ joinˡ⁻ (suc _) k₂ (1# , t₁) t₃ bal = (1# , node k₂ t₁ t₃ bal)
+ joinˡ⁻ n k₂ (## , t₁) t₃ bal
+
+ joinʳ⁻ : ∀ {l u hˡ} hʳ {h} →
+ (k : K& V) →
+ Tree V l [ proj₁ k ] hˡ →
+ (∃ λ i → Tree V [ proj₁ k ] u pred[ i ⊕ hʳ ]) →
+ (bal : hˡ ∼ hʳ ⊔ h) →
+ ∃ λ i → Tree V l u (i ⊕ h)
+ joinʳ⁻ zero k₂ t₁ (0# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
+ joinʳ⁻ zero k₂ t₁ (1# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
+ joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼- = joinˡ⁺ k₂ (1# , t₁) t₃ ∼-
+ joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼0 = (1# , node k₂ t₁ t₃ ∼-)
+ joinʳ⁻ (suc _) k₂ t₁ (0# , t₃) ∼+ = (0# , node k₂ t₁ t₃ ∼0)
+ joinʳ⁻ (suc _) k₂ t₁ (1# , t₃) bal = (1# , node k₂ t₁ t₃ bal)
+ joinʳ⁻ n k₂ t₁ (## , t₃) bal
+
+ -- Extracts the smallest element from the tree, plus the rest.
+ -- Logarithmic in the size of the tree.
+
+ headTail : ∀ {l u h} → Tree V l u (1 + h) →
+ ∃ λ (k : K& V) → l <⁺ [ proj₁ k ] ×
+ ∃ λ i → Tree V [ proj₁ k ] u (i ⊕ h)
+ headTail (node k₁ (leaf l<k₁) t₂ ∼+) = (k₁ , l<k₁ , 0# , t₂)
+ headTail (node k₁ (leaf l<k₁) t₂ ∼0) = (k₁ , l<k₁ , 0# , t₂)
+ headTail (node {hˡ = suc _} k₃ t₁₂ t₄ bal) with headTail t₁₂
+ ... | (k₁ , l<k₁ , t₂) = (k₁ , l<k₁ , joinˡ⁻ _ k₃ t₂ t₄ bal)
+
+ -- Extracts the largest element from the tree, plus the rest.
+ -- Logarithmic in the size of the tree.
+
+ initLast : ∀ {l u h} → Tree V l u (1 + h) →
+ ∃ λ (k : K& V) → [ proj₁ k ] <⁺ u ×
+ ∃ λ i → Tree V l [ proj₁ k ] (i ⊕ h)
+ initLast (node k₂ t₁ (leaf k₂<u) ∼-) = (k₂ , k₂<u , (0# , t₁))
+ initLast (node k₂ t₁ (leaf k₂<u) ∼0) = (k₂ , k₂<u , (0# , t₁))
+ initLast (node {hʳ = suc _} k₂ t₁ t₃₄ bal) with initLast t₃₄
+ ... | (k₄ , k₄<u , t₃) = (k₄ , k₄<u , joinʳ⁻ _ k₂ t₁ t₃ bal)
+
+ -- Another joining function. Logarithmic in the size of either of
+ -- the input trees (which need to have almost equal heights).
+
+ join : ∀ {l m u hˡ hʳ h} →
+ Tree V l m hˡ → Tree V m u hʳ → (bal : hˡ ∼ hʳ ⊔ h) →
+ ∃ λ i → Tree V l u (i ⊕ h)
+ join t₁ (leaf m<u) ∼0 = (0# , castʳ t₁ m<u)
+ join t₁ (leaf m<u) ∼- = (0# , castʳ t₁ m<u)
+ join {hʳ = suc _} t₁ t₂₃ bal with headTail t₂₃
+ ... | (k₂ , m<k₂ , t₃) = joinʳ⁻ _ k₂ (castʳ t₁ m<k₂) t₃ bal
+
+ -- An empty tree.
+
+ empty : ∀ {l u} → l <⁺ u → Tree V l u 0
+ empty = leaf
+
+ -- A singleton tree.
+
+ singleton : ∀ {l u} (k : Key) → V k → l < k < u → Tree V l u 1
+ singleton k v (l<k , k<u) = node (k , v) (leaf l<k) (leaf k<u) ∼0
+
+ -- Inserts a key into the tree, using a function to combine any
+ -- existing value with the new value. Logarithmic in the size of the
+ -- tree (assuming constant-time comparisons and a constant-time
+ -- combining function).
+
+ insertWith : ∀ {l u h} → (k : Key) → V k →
+ (V k → V k → V k) → -- New → old → result.
+ Tree V l u h → l < k < u →
+ ∃ λ i → Tree V l u (i ⊕ h)
+ insertWith k v f (leaf l<u) l<k<u = (1# , singleton k v l<k<u)
+ insertWith k v f (node (k′ , v′) lp pu bal) (l<k , k<u) with compare k k′
+ ... | tri< k<k′ _ _ = joinˡ⁺ (k′ , v′) (insertWith k v f lp (l<k , k<k′)) pu bal
+ ... | tri> _ _ k′<k = joinʳ⁺ (k′ , v′) lp (insertWith k v f pu (k′<k , k<u)) bal
+ ... | tri≈ _ k≡k′ _ rewrite P.sym k≡k′ = (0# , node (k , f v v′) lp pu bal)
+
+ -- Inserts a key into the tree. If the key already exists, then it
+ -- is replaced. Logarithmic in the size of the tree (assuming
+ -- constant-time comparisons).
+
+ insert : ∀ {l u h} → (k : Key) → V k → Tree V l u h → l < k < u →
+ ∃ λ i → Tree V l u (i ⊕ h)
+ insert k v = insertWith k v const
+
+ -- Deletes the key/value pair containing the given key, if any.
+ -- Logarithmic in the size of the tree (assuming constant-time
+ -- comparisons).
+
+ delete : ∀ {l u h} → Key → Tree V l u h →
+ ∃ λ i → Tree V l u pred[ i ⊕ h ]
+ delete k (leaf l<u) = (0# , leaf l<u)
+ delete k (node p lp pu bal) with compare k (proj₁ p)
+ ... | tri< _ _ _ = joinˡ⁻ _ p (delete k lp) pu bal
+ ... | tri> _ _ _ = joinʳ⁻ _ p lp (delete k pu) bal
+ ... | tri≈ _ _ _ = join lp pu bal
+
+ -- Looks up a key. Logarithmic in the size of the tree (assuming
+ -- constant-time comparisons).
+
+ lookup : ∀ {l u h} → (k : Key) → Tree V l u h → Maybe (V k)
+ lookup k (leaf _) = nothing
+ lookup k (node (k′ , v) lk′ k′u _) with compare k k′
+ ... | tri< _ _ _ = lookup k lk′
+ ... | tri> _ _ _ = lookup k k′u
+ ... | tri≈ _ eq _ rewrite eq = just v
+
+ -- Converts the tree to an ordered list. Linear in the size of the
+ -- tree.
+
+ open DiffList
+
+ toDiffList : ∀ {l u h} → Tree V l u h → DiffList (K& V)
+ toDiffList (leaf _) = []
+ toDiffList (node k l r _) = toDiffList l ++ k ∷ toDiffList r
+
+module _ {v w} {V : Key → Set v} {W : Key → Set w} where
+
+ -- Maps a function over all values in the tree.
+
+ map : (∀ {k} → V k → W k) → ∀ {l u h} → Tree V l u h → Tree W l u h
+ map f (leaf l<u) = leaf l<u
+ map f (node (k , v) l r bal) = node (k , f v) (map f l) (map f r) bal
+
diff --git a/src/Data/AVL/IndexedMap.agda b/src/Data/AVL/IndexedMap.agda
index b0e1401..b2bf3a0 100644
--- a/src/Data/AVL/IndexedMap.agda
+++ b/src/Data/AVL/IndexedMap.agda
@@ -41,8 +41,9 @@ private
private
open module AVL =
- Data.AVL (λ ik → Value (proj₁ ik)) isStrictTotalOrder
- public using () renaming (Tree to Map)
+ Data.AVL isStrictTotalOrder
+ public using () renaming (Tree to Map')
+ Map = Map' (Value ∘ proj₁)
-- Repackaged functions.
@@ -50,21 +51,21 @@ empty : Map
empty = AVL.empty
singleton : ∀ {i} → Key i → Value i → Map
-singleton k v = AVL.singleton (, k) v
+singleton k v = AVL.singleton (-, k) v
insert : ∀ {i} → Key i → Value i → Map → Map
-insert k v = AVL.insert (, k) v
+insert k v = AVL.insert (-, k) v
delete : ∀ {i} → Key i → Map → Map
-delete k = AVL.delete (, k)
+delete k = AVL.delete (-, k)
lookup : ∀ {i} → Key i → Map → Maybe (Value i)
-lookup k m = AVL.lookup (, k) m
+lookup k m = AVL.lookup (-, k) m
infix 4 _∈?_
_∈?_ : ∀ {i} → Key i → Map → Bool
-_∈?_ k = AVL._∈?_ (, k)
+_∈?_ k = AVL._∈?_ (-, k)
headTail : Map → Maybe (KV × Map)
headTail m = Maybe.map (Prod.map toKV id) (AVL.headTail m)
diff --git a/src/Data/AVL/Key.agda b/src/Data/AVL/Key.agda
new file mode 100644
index 0000000..78dd689
--- /dev/null
+++ b/src/Data/AVL/Key.agda
@@ -0,0 +1,68 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Keys for AVL trees
+-- The key type extended with a new minimum and maximum.
+-----------------------------------------------------------------------
+
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_ ; refl)
+
+module Data.AVL.Key
+ {k r} (Key : Set k)
+ {_<_ : Rel Key r}
+ (isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_)
+ where
+
+open IsStrictTotalOrder isStrictTotalOrder
+
+open import Level
+open import Data.Empty
+open import Data.Unit
+open import Data.Product
+
+infix 5 [_]
+
+data Key⁺ : Set k where
+ ⊥⁺ ⊤⁺ : Key⁺
+ [_] : (k : Key) → Key⁺
+
+[_]-injective : ∀ {k l} → [ k ] ≡ [ l ] → k ≡ l
+[_]-injective refl = refl
+
+-- An extended strict ordering relation.
+
+infix 4 _<⁺_
+
+_<⁺_ : Key⁺ → Key⁺ → Set r
+⊥⁺ <⁺ [ _ ] = Lift r ⊤
+⊥⁺ <⁺ ⊤⁺ = Lift r ⊤
+[ x ] <⁺ [ y ] = x < y
+[ _ ] <⁺ ⊤⁺ = Lift r ⊤
+_ <⁺ _ = Lift r ⊥
+
+-- A pair of ordering constraints.
+
+infix 4 _<_<_
+
+_<_<_ : Key⁺ → Key → Key⁺ → Set r
+l < x < u = l <⁺ [ x ] × [ x ] <⁺ u
+
+-- _<⁺_ is transitive.
+
+trans⁺ : ∀ l {m u} → l <⁺ m → m <⁺ u → l <⁺ u
+
+trans⁺ [ l ] {m = [ m ]} {u = [ u ]} l<m m<u = trans l<m m<u
+
+trans⁺ ⊥⁺ {u = [ _ ]} _ _ = _
+trans⁺ ⊥⁺ {u = ⊤⁺} _ _ = _
+trans⁺ [ _ ] {u = ⊤⁺} _ _ = _
+
+trans⁺ _ {m = ⊥⁺} {u = ⊥⁺} _ (lift ())
+trans⁺ _ {m = [ _ ]} {u = ⊥⁺} _ (lift ())
+trans⁺ _ {m = ⊤⁺} {u = ⊥⁺} _ (lift ())
+trans⁺ [ _ ] {m = ⊥⁺} {u = [ _ ]} (lift ()) _
+trans⁺ [ _ ] {m = ⊤⁺} {u = [ _ ]} _ (lift ())
+trans⁺ ⊤⁺ {m = ⊥⁺} (lift ()) _
+trans⁺ ⊤⁺ {m = [ _ ]} (lift ()) _
+trans⁺ ⊤⁺ {m = ⊤⁺} (lift ()) _
diff --git a/src/Data/AVL/Sets.agda b/src/Data/AVL/Sets.agda
index db0db4f..96fbecf 100644
--- a/src/Data/AVL/Sets.agda
+++ b/src/Data/AVL/Sets.agda
@@ -24,8 +24,9 @@ open import Level
-- The set type. (Note that Set is a reserved word.)
private
- open module S = AVL (const ⊤) isStrictTotalOrder
- public using () renaming (Tree to ⟨Set⟩)
+ open module S = AVL isStrictTotalOrder
+ public using () renaming (Tree to ⟨Set⟩')
+ ⟨Set⟩ = ⟨Set⟩' (const ⊤)
-- Repackaged functions.
diff --git a/src/Data/Bin.agda b/src/Data/Bin.agda
index 7807763..ceef2df 100644
--- a/src/Data/Bin.agda
+++ b/src/Data/Bin.agda
@@ -18,7 +18,6 @@ open import Data.Product using (uncurry; _,_; _×_)
open import Relation.Binary
open import Relation.Binary.PropositionalEquality
using (_≡_; _≢_; refl; sym)
-open import Relation.Binary.List.StrictLex
open import Relation.Nullary
open import Relation.Nullary.Decidable
@@ -104,6 +103,9 @@ infix 4 _<_
data _<_ (b₁ b₂ : Bin) : Set where
less : (lt : (Nat._<_ on toℕ) b₁ b₂) → b₁ < b₂
+less-injective : ∀ {b₁ b₂} {lt₁ lt₂} → (b₁ < b₂ ∋ less lt₁) ≡ less lt₂ → lt₁ ≡ lt₂
+less-injective refl = refl
+
------------------------------------------------------------------------
-- Arithmetic
@@ -236,7 +238,7 @@ private
nats = List.downFrom testLimit
nats⁺ : List ℕ
- nats⁺ = filter (λ n → ⌊ 1 Nat.≤? n ⌋) nats
+ nats⁺ = filter (1 Nat.≤?_) nats
natPairs : List (ℕ × ℕ)
natPairs = List.zip nats (reverse nats)
diff --git a/src/Data/Bin/Properties.agda b/src/Data/Bin/Properties.agda
index 9c95f7b..cedfb27 100644
--- a/src/Data/Bin/Properties.agda
+++ b/src/Data/Bin/Properties.agda
@@ -13,7 +13,7 @@ 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)
+ using (ℕ; zero; z≤n; s≤s)
renaming (suc to 1+_; _+_ to _+ℕ_; _*_ to _*ℕ_; _≤_ to _≤ℕ_)
import Data.Nat.Properties as ℕₚ
open import Data.Product using (proj₁; proj₂)
@@ -36,7 +36,7 @@ _≟ₑ_ : ∀ {base} → Decidable (_≡_ {A = Expansion base})
_≟ₑ_ [] [] = yes refl
_≟ₑ_ [] (_ ∷ _) = no λ()
_≟ₑ_ (_ ∷ _) [] = no λ()
-_≟ₑ_ (x ∷ xs) (y ∷ ys) with x 𝔽ₚ.≟ y | xs ≟ₑ ys
+_≟ₑ_ (x ∷ xs) (y ∷ ys) with x Fin.≟ 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
@@ -72,7 +72,7 @@ as 1# ≟ bs 1# with as ≟ₑ bs
∷ʳ-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+ (m₁ +ℕ n₁ *ℕ 2) ≤⟨ s≤s (ℕₚ.+-monoˡ-≤ _ (𝔽ₚ.toℕ≤pred[n] a)) ⟩
1+ (1 +ℕ n₁ *ℕ 2) ≡⟨ refl ⟩
1+ n₁ *ℕ 2 ≤⟨ ℕₚ.*-mono-≤ lt ℕₚ.≤-refl ⟩
n₂ *ℕ 2 ≤⟨ ℕₚ.n≤m+n m₂ (n₂ *ℕ 2) ⟩
@@ -85,7 +85,7 @@ as 1# ≟ bs 1# with as ≟ₑ bs
∷ˡ-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 ⟩
+ (1 +ℕ m₁) +ℕ n *ℕ 2 ≤⟨ ℕₚ.+-monoˡ-≤ _ lt ⟩
m₂ +ℕ n *ℕ 2 ∎)
where
open ℕₚ.≤-Reasoning
@@ -117,7 +117,7 @@ as 1# ≟ bs 1# with as ≟ₑ bs
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 with 𝔽ₚ.<-cmp a b
... | tri≈ ¬lt′ refl ¬gt′ =
tri≈ (<-irrefl refl) refl (<-irrefl refl)
... | tri< lt′ ¬eq ¬gt′ =
diff --git a/src/Data/Bool/Base.agda b/src/Data/Bool/Base.agda
index 2208199..7a88b91 100644
--- a/src/Data/Bool/Base.agda
+++ b/src/Data/Bool/Base.agda
@@ -9,6 +9,7 @@ open import Data.Unit.Base using (⊤)
open import Data.Empty
open import Relation.Nullary
open import Relation.Binary.Core
+open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl)
infixr 6 _∧_
infixr 5 _∨_ _xor_
diff --git a/src/Data/Bool/Properties.agda b/src/Data/Bool/Properties.agda
index 030d170..1b2b53d 100644
--- a/src/Data/Bool/Properties.agda
+++ b/src/Data/Bool/Properties.agda
@@ -6,245 +6,327 @@
module Data.Bool.Properties where
-open import Data.Bool as Bool
+open import Algebra
+open import Data.Bool
+open import Data.Empty
+open import Data.Product
+open import Data.Sum
open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Equivalence
using (_⇔_; equivalence; module Equivalence)
-open import Algebra
-open import Algebra.Structures
-import Algebra.RingSolver.Simple as Solver
-import Algebra.RingSolver.AlmostCommutativeRing as ACR
-open import Relation.Binary.PropositionalEquality as P
- using (_≡_; _≢_; refl)
-open P.≡-Reasoning
-import Algebra.FunctionProperties as FP; open FP (_≡_ {A = Bool})
-open import Data.Product
-open import Data.Sum
-open import Data.Empty
+open import Relation.Binary.PropositionalEquality
+ hiding ([_]; proof-irrelevance)
+open import Relation.Unary using (Irrelevant)
-------------------------------------------------------------------------
--- Duality
-
--- Can we take advantage of duality in some (nice) way?
+open import Algebra.FunctionProperties (_≡_ {A = Bool})
+open import Algebra.Structures (_≡_ {A = Bool})
+open ≡-Reasoning
------------------------------------------------------------------------
--- (Bool, ∨, ∧, false, true) forms a commutative semiring
+-- Properties of _∨_
∨-assoc : Associative _∨_
∨-assoc true y z = refl
∨-assoc false y z = refl
-∧-assoc : Associative _∧_
-∧-assoc true y z = refl
-∧-assoc false y z = refl
-
∨-comm : Commutative _∨_
∨-comm true true = refl
∨-comm true false = refl
∨-comm false true = refl
∨-comm false false = refl
+∨-identityˡ : LeftIdentity false _∨_
+∨-identityˡ _ = refl
+
+∨-identityʳ : RightIdentity false _∨_
+∨-identityʳ false = refl
+∨-identityʳ true = refl
+
+∨-identity : Identity false _∨_
+∨-identity = ∨-identityˡ , ∨-identityʳ
+
+∨-zeroˡ : LeftZero true _∨_
+∨-zeroˡ _ = refl
+
+∨-zeroʳ : RightZero true _∨_
+∨-zeroʳ false = refl
+∨-zeroʳ true = refl
+
+∨-zero : Zero true _∨_
+∨-zero = ∨-zeroˡ , ∨-zeroʳ
+
+∨-inverseˡ : LeftInverse true not _∨_
+∨-inverseˡ false = refl
+∨-inverseˡ true = refl
+
+∨-inverseʳ : RightInverse true not _∨_
+∨-inverseʳ x = ∨-comm x (not x) ⟨ trans ⟩ ∨-inverseˡ x
+
+∨-inverse : Inverse true not _∨_
+∨-inverse = ∨-inverseˡ , ∨-inverseʳ
+
+∨-idem : Idempotent _∨_
+∨-idem false = refl
+∨-idem true = refl
+
+∨-sel : Selective _∨_
+∨-sel false y = inj₂ refl
+∨-sel true y = inj₁ refl
+
+∨-isSemigroup : IsSemigroup _∨_
+∨-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∨-assoc
+ ; ∙-cong = cong₂ _∨_
+ }
+
+∨-semigroup : Semigroup _ _
+∨-semigroup = record
+ { isSemigroup = ∨-isSemigroup
+ }
+
+∨-isCommutativeMonoid : IsCommutativeMonoid _∨_ false
+∨-isCommutativeMonoid = record
+ { isSemigroup = ∨-isSemigroup
+ ; identityˡ = ∨-identityˡ
+ ; comm = ∨-comm
+ }
+
+∨-commutativeMonoid : CommutativeMonoid _ _
+∨-commutativeMonoid = record
+ { isCommutativeMonoid = ∨-isCommutativeMonoid
+ }
+
+∨-isIdempotentCommutativeMonoid :
+ IsIdempotentCommutativeMonoid _∨_ false
+∨-isIdempotentCommutativeMonoid = record
+ { isCommutativeMonoid = ∨-isCommutativeMonoid
+ ; idem = ∨-idem
+ }
+
+∨-idempotentCommutativeMonoid : IdempotentCommutativeMonoid _ _
+∨-idempotentCommutativeMonoid = record
+ { isIdempotentCommutativeMonoid = ∨-isIdempotentCommutativeMonoid
+ }
+
+------------------------------------------------------------------------
+-- Properties of _∧_
+
+∧-assoc : Associative _∧_
+∧-assoc true y z = refl
+∧-assoc false y z = refl
+
∧-comm : Commutative _∧_
∧-comm true true = refl
∧-comm true false = refl
∧-comm false true = refl
∧-comm false false = refl
-∧-∨-distˡ : _∧_ DistributesOverˡ _∨_
-∧-∨-distˡ true y z = refl
-∧-∨-distˡ false y z = refl
-
-∧-∨-distʳ : _∧_ DistributesOverʳ _∨_
-∧-∨-distʳ x y z =
- begin
- (y ∨ z) ∧ x
- ≡⟨ ∧-comm (y ∨ z) x ⟩
- x ∧ (y ∨ z)
- ≡⟨ ∧-∨-distˡ x y z ⟩
- x ∧ y ∨ x ∧ z
- ≡⟨ P.cong₂ _∨_ (∧-comm x y) (∧-comm x z) ⟩
- y ∧ x ∨ z ∧ x
- ∎
-
-distrib-∧-∨ : _∧_ DistributesOver _∨_
-distrib-∧-∨ = ∧-∨-distˡ , ∧-∨-distʳ
-
-isCommutativeSemiring-∨-∧
- : IsCommutativeSemiring _≡_ _∨_ _∧_ false true
-isCommutativeSemiring-∨-∧ = record
- { +-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = P.isEquivalence
- ; assoc = ∨-assoc
- ; ∙-cong = P.cong₂ _∨_
- }
- ; identityˡ = λ _ → refl
- ; comm = ∨-comm
- }
- ; *-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = P.isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = P.cong₂ _∧_
- }
- ; identityˡ = λ _ → refl
- ; comm = ∧-comm
- }
- ; distribʳ = proj₂ distrib-∧-∨
- ; zeroˡ = λ _ → refl
+∧-identityˡ : LeftIdentity true _∧_
+∧-identityˡ _ = refl
+
+∧-identityʳ : RightIdentity true _∧_
+∧-identityʳ false = refl
+∧-identityʳ true = refl
+
+∧-identity : Identity true _∧_
+∧-identity = ∧-identityˡ , ∧-identityʳ
+
+∧-zeroˡ : LeftZero false _∧_
+∧-zeroˡ _ = refl
+
+∧-zeroʳ : RightZero false _∧_
+∧-zeroʳ false = refl
+∧-zeroʳ true = refl
+
+∧-zero : Zero false _∧_
+∧-zero = ∧-zeroˡ , ∧-zeroʳ
+
+∧-inverseˡ : LeftInverse false not _∧_
+∧-inverseˡ false = refl
+∧-inverseˡ true = refl
+
+∧-inverseʳ : RightInverse false not _∧_
+∧-inverseʳ x = ∧-comm x (not x) ⟨ trans ⟩ ∧-inverseˡ x
+
+∧-inverse : Inverse false not _∧_
+∧-inverse = ∧-inverseˡ , ∧-inverseʳ
+
+∧-idem : Idempotent _∧_
+∧-idem false = refl
+∧-idem true = refl
+
+∧-sel : Selective _∧_
+∧-sel false y = inj₁ refl
+∧-sel true y = inj₂ refl
+
+∧-distribˡ-∨ : _∧_ DistributesOverˡ _∨_
+∧-distribˡ-∨ true y z = refl
+∧-distribˡ-∨ false y z = refl
+
+∧-distribʳ-∨ : _∧_ DistributesOverʳ _∨_
+∧-distribʳ-∨ x y z = begin
+ (y ∨ z) ∧ x ≡⟨ ∧-comm (y ∨ z) x ⟩
+ x ∧ (y ∨ z) ≡⟨ ∧-distribˡ-∨ x y z ⟩
+ x ∧ y ∨ x ∧ z ≡⟨ cong₂ _∨_ (∧-comm x y) (∧-comm x z) ⟩
+ y ∧ x ∨ z ∧ x ∎
+
+∧-distrib-∨ : _∧_ DistributesOver _∨_
+∧-distrib-∨ = ∧-distribˡ-∨ , ∧-distribʳ-∨
+
+∨-distribˡ-∧ : _∨_ DistributesOverˡ _∧_
+∨-distribˡ-∧ true y z = refl
+∨-distribˡ-∧ false y z = refl
+
+∨-distribʳ-∧ : _∨_ DistributesOverʳ _∧_
+∨-distribʳ-∧ x y z = begin
+ (y ∧ z) ∨ x ≡⟨ ∨-comm (y ∧ z) x ⟩
+ x ∨ (y ∧ z) ≡⟨ ∨-distribˡ-∧ x y z ⟩
+ (x ∨ y) ∧ (x ∨ z) ≡⟨ cong₂ _∧_ (∨-comm x y) (∨-comm x z) ⟩
+ (y ∨ x) ∧ (z ∨ x) ∎
+
+∨-distrib-∧ : _∨_ DistributesOver _∧_
+∨-distrib-∧ = ∨-distribˡ-∧ , ∨-distribʳ-∧
+
+∧-abs-∨ : _∧_ Absorbs _∨_
+∧-abs-∨ true y = refl
+∧-abs-∨ false y = refl
+
+∨-abs-∧ : _∨_ Absorbs _∧_
+∨-abs-∧ true y = refl
+∨-abs-∧ false y = refl
+
+∨-∧-absorptive : Absorptive _∨_ _∧_
+∨-∧-absorptive = ∨-abs-∧ , ∧-abs-∨
+
+∧-isSemigroup : IsSemigroup _∧_
+∧-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∧-assoc
+ ; ∙-cong = cong₂ _∧_
+ }
+
+∧-semigroup : Semigroup _ _
+∧-semigroup = record
+ { isSemigroup = ∧-isSemigroup
+ }
+
+∧-isCommutativeMonoid : IsCommutativeMonoid _∧_ true
+∧-isCommutativeMonoid = record
+ { isSemigroup = ∧-isSemigroup
+ ; identityˡ = ∧-identityˡ
+ ; comm = ∧-comm
+ }
+
+∧-commutativeMonoid : CommutativeMonoid _ _
+∧-commutativeMonoid = record
+ { isCommutativeMonoid = ∧-isCommutativeMonoid
}
-commutativeSemiring-∨-∧ : CommutativeSemiring _ _
-commutativeSemiring-∨-∧ = record
+∧-isIdempotentCommutativeMonoid :
+ IsIdempotentCommutativeMonoid _∧_ true
+∧-isIdempotentCommutativeMonoid = record
+ { isCommutativeMonoid = ∧-isCommutativeMonoid
+ ; idem = ∧-idem
+ }
+
+∧-idempotentCommutativeMonoid : IdempotentCommutativeMonoid _ _
+∧-idempotentCommutativeMonoid = record
+ { isIdempotentCommutativeMonoid = ∧-isIdempotentCommutativeMonoid
+ }
+
+∨-∧-isCommutativeSemiring
+ : IsCommutativeSemiring _∨_ _∧_ false true
+∨-∧-isCommutativeSemiring = record
+ { +-isCommutativeMonoid = ∨-isCommutativeMonoid
+ ; *-isCommutativeMonoid = ∧-isCommutativeMonoid
+ ; distribʳ = ∧-distribʳ-∨
+ ; zeroˡ = ∧-zeroˡ
+ }
+
+∨-∧-commutativeSemiring : CommutativeSemiring _ _
+∨-∧-commutativeSemiring = record
{ _+_ = _∨_
; _*_ = _∧_
; 0# = false
; 1# = true
- ; isCommutativeSemiring = isCommutativeSemiring-∨-∧
+ ; isCommutativeSemiring = ∨-∧-isCommutativeSemiring
}
-module RingSolver =
- Solver (ACR.fromCommutativeSemiring commutativeSemiring-∨-∧) _≟_
-
-------------------------------------------------------------------------
--- (Bool, ∧, ∨, true, false) forms a commutative semiring
-
-∨-∧-distˡ : _∨_ DistributesOverˡ _∧_
-∨-∧-distˡ true y z = refl
-∨-∧-distˡ false y z = refl
-
-∨-∧-distʳ : _∨_ DistributesOverʳ _∧_
-∨-∧-distʳ x y z =
- begin
- (y ∧ z) ∨ x
- ≡⟨ ∨-comm (y ∧ z) x ⟩
- x ∨ (y ∧ z)
- ≡⟨ ∨-∧-distˡ x y z ⟩
- (x ∨ y) ∧ (x ∨ z)
- ≡⟨ P.cong₂ _∧_ (∨-comm x y) (∨-comm x z) ⟩
- (y ∨ x) ∧ (z ∨ x)
- ∎
-
-∨-∧-distrib : _∨_ DistributesOver _∧_
-∨-∧-distrib = ∨-∧-distˡ , ∨-∧-distʳ
-
-isCommutativeSemiring-∧-∨
- : IsCommutativeSemiring _≡_ _∧_ _∨_ true false
-isCommutativeSemiring-∧-∨ = record
- { +-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = P.isEquivalence
- ; assoc = ∧-assoc
- ; ∙-cong = P.cong₂ _∧_
- }
- ; identityˡ = λ _ → refl
- ; comm = ∧-comm
- }
- ; *-isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = P.isEquivalence
- ; assoc = ∨-assoc
- ; ∙-cong = P.cong₂ _∨_
- }
- ; identityˡ = λ _ → refl
- ; comm = ∨-comm
- }
- ; distribʳ = ∨-∧-distʳ
- ; zeroˡ = λ _ → refl
+∧-∨-isCommutativeSemiring
+ : IsCommutativeSemiring _∧_ _∨_ true false
+∧-∨-isCommutativeSemiring = record
+ { +-isCommutativeMonoid = ∧-isCommutativeMonoid
+ ; *-isCommutativeMonoid = ∨-isCommutativeMonoid
+ ; distribʳ = ∨-distribʳ-∧
+ ; zeroˡ = ∨-zeroˡ
}
-commutativeSemiring-∧-∨ : CommutativeSemiring _ _
-commutativeSemiring-∧-∨ = record
+∧-∨-commutativeSemiring : CommutativeSemiring _ _
+∧-∨-commutativeSemiring = record
{ _+_ = _∧_
; _*_ = _∨_
; 0# = true
; 1# = false
- ; isCommutativeSemiring = isCommutativeSemiring-∧-∨
+ ; isCommutativeSemiring = ∧-∨-isCommutativeSemiring
}
-------------------------------------------------------------------------
--- (Bool, ∨, ∧, not, true, false) is a boolean algebra
+∨-∧-isLattice : IsLattice _∨_ _∧_
+∨-∧-isLattice = record
+ { isEquivalence = isEquivalence
+ ; ∨-comm = ∨-comm
+ ; ∨-assoc = ∨-assoc
+ ; ∨-cong = cong₂ _∨_
+ ; ∧-comm = ∧-comm
+ ; ∧-assoc = ∧-assoc
+ ; ∧-cong = cong₂ _∧_
+ ; absorptive = ∨-∧-absorptive
+ }
-∨-∧-abs : _∨_ Absorbs _∧_
-∨-∧-abs true y = refl
-∨-∧-abs false y = refl
+∨-∧-lattice : Lattice _ _
+∨-∧-lattice = record
+ { isLattice = ∨-∧-isLattice
+ }
-∧-∨-abs : _∧_ Absorbs _∨_
-∧-∨-abs true y = refl
-∧-∨-abs false y = refl
+∨-∧-isDistributiveLattice : IsDistributiveLattice _∨_ _∧_
+∨-∧-isDistributiveLattice = record
+ { isLattice = ∨-∧-isLattice
+ ; ∨-∧-distribʳ = ∨-distribʳ-∧
+ }
-∨-∧-absorptive : Absorptive _∨_ _∧_
-∨-∧-absorptive = ∨-∧-abs , ∧-∨-abs
-
-not-∧-inverseˡ : LeftInverse false not _∧_
-not-∧-inverseˡ false = refl
-not-∧-inverseˡ true = refl
-
-not-∧-inverseʳ : RightInverse false not _∧_
-not-∧-inverseʳ x = ∧-comm x (not x) ⟨ P.trans ⟩ not-∧-inverseˡ x
-
-not-∧-inverse : Inverse false not _∧_
-not-∧-inverse = not-∧-inverseˡ , not-∧-inverseʳ
-
-not-∨-inverseˡ : LeftInverse true not _∨_
-not-∨-inverseˡ false = refl
-not-∨-inverseˡ true = refl
-
-not-∨-inverseʳ : RightInverse true not _∨_
-not-∨-inverseʳ x = ∨-comm x (not x) ⟨ P.trans ⟩ not-∨-inverseˡ x
-
-not-∨-inverse : Inverse true not _∨_
-not-∨-inverse = not-∨-inverseˡ , not-∨-inverseʳ
-
-isBooleanAlgebra : IsBooleanAlgebra _≡_ _∨_ _∧_ not true false
-isBooleanAlgebra = record
- { isDistributiveLattice = record
- { isLattice = record
- { isEquivalence = P.isEquivalence
- ; ∨-comm = ∨-comm
- ; ∨-assoc = ∨-assoc
- ; ∨-cong = P.cong₂ _∨_
- ; ∧-comm = ∧-comm
- ; ∧-assoc = ∧-assoc
- ; ∧-cong = P.cong₂ _∧_
- ; absorptive = ∨-∧-absorptive
- }
- ; ∨-∧-distribʳ = ∨-∧-distʳ
- }
- ; ∨-complementʳ = not-∨-inverseʳ
- ; ∧-complementʳ = not-∧-inverseʳ
- ; ¬-cong = P.cong not
+∨-∧-distributiveLattice : DistributiveLattice _ _
+∨-∧-distributiveLattice = record
+ { isDistributiveLattice = ∨-∧-isDistributiveLattice
}
-booleanAlgebra : BooleanAlgebra _ _
-booleanAlgebra = record
- { _∨_ = _∨_
- ; _∧_ = _∧_
- ; ¬_ = not
- ; ⊤ = true
- ; ⊥ = false
- ; isBooleanAlgebra = isBooleanAlgebra
+∨-∧-isBooleanAlgebra : IsBooleanAlgebra _∨_ _∧_ not true false
+∨-∧-isBooleanAlgebra = record
+ { isDistributiveLattice = ∨-∧-isDistributiveLattice
+ ; ∨-complementʳ = ∨-inverseʳ
+ ; ∧-complementʳ = ∧-inverseʳ
+ ; ¬-cong = cong not
+ }
+
+∨-∧-booleanAlgebra : BooleanAlgebra _ _
+∨-∧-booleanAlgebra = record
+ { isBooleanAlgebra = ∨-∧-isBooleanAlgebra
}
------------------------------------------------------------------------
--- (Bool, xor, ∧, id, false, true) forms a commutative ring
+-- Properties of _xor_
xor-is-ok : ∀ x y → x xor y ≡ (x ∨ y) ∧ not (x ∧ y)
xor-is-ok true y = refl
-xor-is-ok false y = P.sym $ proj₂ CS.*-identity _
- where module CS = CommutativeSemiring commutativeSemiring-∨-∧
+xor-is-ok false y = sym (∧-identityʳ _)
-commutativeRing-xor-∧ : CommutativeRing _ _
-commutativeRing-xor-∧ = commutativeRing
+xor-∧-commutativeRing : CommutativeRing _ _
+xor-∧-commutativeRing = commutativeRing
where
import Algebra.Properties.BooleanAlgebra as BA
- open BA booleanAlgebra
+ open BA ∨-∧-booleanAlgebra
open XorRing _xor_ xor-is-ok
-module XorRingSolver =
- Solver (ACR.fromCommutativeRing commutativeRing-xor-∧) _≟_
-
------------------------------------------------------------------------
-- Miscellaneous other properties
@@ -264,10 +346,10 @@ not-¬ {false} refl ()
⇔→≡ : {b₁ b₂ b : Bool} → b₁ ≡ b ⇔ b₂ ≡ b → b₁ ≡ b₂
⇔→≡ {true } {true } hyp = refl
-⇔→≡ {true } {false} {true } hyp = P.sym (Equivalence.to hyp ⟨$⟩ refl)
+⇔→≡ {true } {false} {true } hyp = sym (Equivalence.to hyp ⟨$⟩ refl)
⇔→≡ {true } {false} {false} hyp = Equivalence.from hyp ⟨$⟩ refl
⇔→≡ {false} {true } {true } hyp = Equivalence.from hyp ⟨$⟩ refl
-⇔→≡ {false} {true } {false} hyp = P.sym (Equivalence.to hyp ⟨$⟩ refl)
+⇔→≡ {false} {true } {false} hyp = sym (Equivalence.to hyp ⟨$⟩ refl)
⇔→≡ {false} {false} hyp = refl
T-≡ : ∀ {b} → T b ⇔ b ≡ true
@@ -288,12 +370,131 @@ T-∨ {true} {b₂} = equivalence inj₁ (const _)
T-∨ {false} {true} = equivalence inj₂ (const _)
T-∨ {false} {false} = equivalence inj₁ [ id , id ]
-proof-irrelevance : ∀ {b} (p q : T b) → p ≡ q
-proof-irrelevance {true} _ _ = refl
-proof-irrelevance {false} () ()
+T-irrelevance : Irrelevant T
+T-irrelevance {true} _ _ = refl
+T-irrelevance {false} () ()
push-function-into-if :
∀ {a b} {A : Set a} {B : Set b} (f : A → B) x {y z} →
f (if x then y else z) ≡ (if x then f y else f z)
-push-function-into-if _ true = P.refl
-push-function-into-if _ false = P.refl
+push-function-into-if _ true = refl
+push-function-into-if _ false = refl
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+∧-∨-distˡ = ∧-distribˡ-∨
+{-# WARNING_ON_USAGE ∧-∨-distˡ
+"Warning: ∧-∨-distˡ was deprecated in v0.15.
+Please use ∧-distribˡ-∨ instead."
+#-}
+∧-∨-distʳ = ∧-distribʳ-∨
+{-# WARNING_ON_USAGE ∧-∨-distʳ
+"Warning: ∧-∨-distʳ was deprecated in v0.15.
+Please use ∧-distribʳ-∨ instead."
+#-}
+distrib-∧-∨ = ∧-distrib-∨
+{-# WARNING_ON_USAGE distrib-∧-∨
+"Warning: distrib-∧-∨ was deprecated in v0.15.
+Please use ∧-distrib-∨ instead."
+#-}
+∨-∧-distˡ = ∨-distribˡ-∧
+{-# WARNING_ON_USAGE ∨-∧-distˡ
+"Warning: ∨-∧-distˡ was deprecated in v0.15.
+Please use ∨-distribˡ-∧ instead."
+#-}
+∨-∧-distʳ = ∨-distribʳ-∧
+{-# WARNING_ON_USAGE ∨-∧-distʳ
+"Warning: ∨-∧-distʳ was deprecated in v0.15.
+Please use ∨-distribʳ-∧ instead."
+#-}
+∨-∧-distrib = ∨-distrib-∧
+{-# WARNING_ON_USAGE ∨-∧-distrib
+"Warning: ∨-∧-distrib was deprecated in v0.15.
+Please use ∨-distrib-∧ instead."
+#-}
+∨-∧-abs = ∨-abs-∧
+{-# WARNING_ON_USAGE ∨-∧-abs
+"Warning: ∨-∧-abs was deprecated in v0.15.
+Please use ∨-abs-∧ instead."
+#-}
+∧-∨-abs = ∧-abs-∨
+{-# WARNING_ON_USAGE ∧-∨-abs
+"Warning: ∧-∨-abs was deprecated in v0.15.
+Please use ∧-abs-∨ instead."
+#-}
+not-∧-inverseˡ = ∧-inverseˡ
+{-# WARNING_ON_USAGE not-∧-inverseˡ
+"Warning: not-∧-inverseˡ was deprecated in v0.15.
+Please use ∧-inverseˡ instead."
+#-}
+not-∧-inverseʳ = ∧-inverseʳ
+{-# WARNING_ON_USAGE not-∧-inverseʳ
+"Warning: not-∧-inverseʳ was deprecated in v0.15.
+Please use ∧-inverseʳ instead."
+#-}
+not-∧-inverse = ∧-inverse
+{-# WARNING_ON_USAGE not-∧-inverse
+"Warning: not-∧-inverse was deprecated in v0.15.
+Please use ∧-inverse instead."
+#-}
+not-∨-inverseˡ = ∨-inverseˡ
+{-# WARNING_ON_USAGE not-∨-inverseˡ
+"Warning: not-∨-inverseˡ was deprecated in v0.15.
+Please use ∨-inverseˡ instead."
+#-}
+not-∨-inverseʳ = ∨-inverseʳ
+{-# WARNING_ON_USAGE not-∨-inverseʳ
+"Warning: not-∨-inverseʳ was deprecated in v0.15.
+Please use ∨-inverseʳ instead."
+#-}
+not-∨-inverse = ∨-inverse
+{-# WARNING_ON_USAGE not-∨-inverse
+"Warning: not-∨-inverse was deprecated in v0.15.
+Please use ∨-inverse instead."
+#-}
+isCommutativeSemiring-∨-∧ = ∨-∧-isCommutativeSemiring
+{-# WARNING_ON_USAGE isCommutativeSemiring-∨-∧
+"Warning: isCommutativeSemiring-∨-∧ was deprecated in v0.15.
+Please use ∨-∧-isCommutativeSemiring instead."
+#-}
+commutativeSemiring-∨-∧ = ∨-∧-commutativeSemiring
+{-# WARNING_ON_USAGE commutativeSemiring-∨-∧
+"Warning: commutativeSemiring-∨-∧ was deprecated in v0.15.
+Please use ∨-∧-commutativeSemiring instead."
+#-}
+isCommutativeSemiring-∧-∨ = ∧-∨-isCommutativeSemiring
+{-# WARNING_ON_USAGE isCommutativeSemiring-∧-∨
+"Warning: isCommutativeSemiring-∧-∨ was deprecated in v0.15.
+Please use ∧-∨-isCommutativeSemiring instead."
+#-}
+commutativeSemiring-∧-∨ = ∧-∨-commutativeSemiring
+{-# WARNING_ON_USAGE commutativeSemiring-∧-∨
+"Warning: commutativeSemiring-∧-∨ was deprecated in v0.15.
+Please use ∧-∨-commutativeSemiring instead."
+#-}
+isBooleanAlgebra = ∨-∧-isBooleanAlgebra
+{-# WARNING_ON_USAGE isBooleanAlgebra
+"Warning: isBooleanAlgebra was deprecated in v0.15.
+Please use ∨-∧-isBooleanAlgebra instead."
+#-}
+booleanAlgebra = ∨-∧-booleanAlgebra
+{-# WARNING_ON_USAGE booleanAlgebra
+"Warning: booleanAlgebra was deprecated in v0.15.
+Please use ∨-∧-booleanAlgebra instead."
+#-}
+commutativeRing-xor-∧ = xor-∧-commutativeRing
+{-# WARNING_ON_USAGE commutativeRing-xor-∧
+"Warning: commutativeRing-xor-∧ was deprecated in v0.15.
+Please use xor-∧-commutativeRing instead."
+#-}
+proof-irrelevance = T-irrelevance
+{-# WARNING_ON_USAGE proof-irrelevance
+"Warning: proof-irrelevance was deprecated in v0.15.
+Please use T-irrelevance instead."
+#-}
diff --git a/src/Data/Bool/Solver.agda b/src/Data/Bool/Solver.agda
new file mode 100644
index 0000000..b3aeedb
--- /dev/null
+++ b/src/Data/Bool/Solver.agda
@@ -0,0 +1,28 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Automatic solvers for equations over booleans
+------------------------------------------------------------------------
+
+-- See README.Nat for examples of how to use similar solvers
+
+module Data.Bool.Solver where
+
+import Algebra.Solver.Ring.Simple as Solver
+import Algebra.Solver.Ring.AlmostCommutativeRing as ACR
+open import Data.Bool using (_≟_)
+open import Data.Bool.Properties
+
+------------------------------------------------------------------------
+-- A module for automatically solving propositional equivalences
+-- containing _∨_ and _∧_
+
+module ∨-∧-Solver =
+ Solver (ACR.fromCommutativeSemiring ∨-∧-commutativeSemiring) _≟_
+
+------------------------------------------------------------------------
+-- A module for automatically solving propositional equivalences
+-- containing _xor_ and _∧_
+
+module xor-∧-Solver =
+ Solver (ACR.fromCommutativeRing xor-∧-commutativeRing) _≟_
diff --git a/src/Data/BoundedVec.agda b/src/Data/BoundedVec.agda
index 6359af2..cd88d6d 100644
--- a/src/Data/BoundedVec.agda
+++ b/src/Data/BoundedVec.agda
@@ -13,8 +13,8 @@ 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
+open import Data.Nat.Solver
+open +-*-Solver
------------------------------------------------------------------------
-- The type
diff --git a/src/Data/Char.agda b/src/Data/Char.agda
index be928a5..aceb9f3 100644
--- a/src/Data/Char.agda
+++ b/src/Data/Char.agda
@@ -6,59 +6,23 @@
module Data.Char where
-open import Data.Nat.Base using (ℕ)
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
-open import Relation.Binary
-import Relation.Binary.On as On
-open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
-open import Relation.Binary.PropositionalEquality.TrustMe
+open import Relation.Binary using (Setoid; StrictTotalOrder)
+import Relation.Binary.Construct.On as On
+import Relation.Binary.PropositionalEquality as PropEq
-open import Data.String.Base using (String)
-open import Data.Char.Base
-open Data.Char.Base public using (Char; show; toNat)
-
--- Informative equality test.
-
-infix 4 _≟_
-
-_≟_ : Decidable {A = Char} _≡_
-s₁ ≟ s₂ with primCharEquality s₁ s₂
-... | true = yes trustMe
-... | false = no whatever
- where postulate whatever : _
-
--- Boolean equality test.
---
--- Why is the definition _==_ = primCharEquality not used? One reason
--- is that the present definition can sometimes improve type
--- inference, at least with the version of Agda that is current at the
--- time of writing: see unit-test below.
-
-infix 4 _==_
-
-_==_ : Char → Char → Bool
-c₁ == c₂ = ⌊ c₁ ≟ c₂ ⌋
-
-private
-
- -- The following unit test does not type-check (at the time of
- -- writing) if _==_ is replaced by primCharEquality.
+------------------------------------------------------------------------
+-- Re-export base definitions publically
- data P : (Char → Bool) → Set where
- p : (c : Char) → P (_==_ c)
+open import Data.Char.Base public
- unit-test : P (_==_ 'x')
- unit-test = p _
+------------------------------------------------------------------------
+-- Equality over characters
setoid : Setoid _ _
setoid = PropEq.setoid Char
-decSetoid : DecSetoid _ _
-decSetoid = PropEq.decSetoid _≟_
-
+------------------------------------------------------------------------
-- An ordering induced by the toNat function.
strictTotalOrder : StrictTotalOrder _ _ _
diff --git a/src/Data/Char/Base.agda b/src/Data/Char/Base.agda
index ba1aa55..a5aa2c8 100644
--- a/src/Data/Char/Base.agda
+++ b/src/Data/Char/Base.agda
@@ -5,23 +5,51 @@
------------------------------------------------------------------------
module Data.Char.Base where
+open import Agda.Builtin.String using (primShowChar)
open import Data.Nat.Base using (ℕ)
open import Data.Bool.Base using (Bool)
open import Data.String.Base using (String)
------------------------------------------------------------------------
--- Re-export the type from the Core module
+-- Re-export the type
-open import Data.Char.Core using (Char) public
+import Agda.Builtin.Char as AgdaChar
+open AgdaChar using (Char) public
------------------------------------------------------------------------
-- Primitive operations
-open import Agda.Builtin.Char public using (primCharToNat; primCharEquality)
-open import Agda.Builtin.String public using (primShowChar)
+open AgdaChar
show : Char → String
show = primShowChar
+isLower : Char → Bool
+isLower = primIsLower
+
+isDigit : Char → Bool
+isDigit = primIsDigit
+
+isAlpha : Char → Bool
+isAlpha = primIsAlpha
+
+isSpace : Char → Bool
+isSpace = primIsSpace
+
+isAscii : Char → Bool
+isAscii = primIsAscii
+
+isLatin1 : Char → Bool
+isLatin1 = primIsLatin1
+
+isPrint : Char → Bool
+isPrint = primIsPrint
+
+isHexDigit : Char → Bool
+isHexDigit = primIsHexDigit
+
toNat : Char → ℕ
toNat = primCharToNat
+
+fromNat : ℕ → Char
+fromNat = primNatToChar
diff --git a/src/Data/Char/Unsafe.agda b/src/Data/Char/Unsafe.agda
new file mode 100644
index 0000000..581afba
--- /dev/null
+++ b/src/Data/Char/Unsafe.agda
@@ -0,0 +1,58 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Unsafe Char operations and proofs
+------------------------------------------------------------------------
+
+module Data.Char.Unsafe where
+
+open import Data.Bool.Base using (Bool; true; false)
+open import Relation.Nullary using (yes; no)
+open import Relation.Nullary.Decidable using (⌊_⌋)
+open import Relation.Binary using (Decidable; DecSetoid)
+open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
+open import Relation.Binary.PropositionalEquality.TrustMe
+
+open import Agda.Builtin.Char using (primCharEquality)
+open import Data.Char
+
+------------------------------------------------------------------------
+-- An informative equality test.
+
+infix 4 _≟_
+
+_≟_ : Decidable {A = Char} _≡_
+s₁ ≟ s₂ with primCharEquality s₁ s₂
+... | true = yes trustMe
+... | false = no whatever
+ where postulate whatever : _
+
+------------------------------------------------------------------------
+-- Boolean equality test.
+--
+-- Why is the definition _==_ = primCharEquality not used? One reason
+-- is that the present definition can sometimes improve type
+-- inference, at least with the version of Agda that is current at the
+-- time of writing: see unit-test below.
+
+infix 4 _==_
+
+_==_ : Char → Char → Bool
+c₁ == c₂ = ⌊ c₁ ≟ c₂ ⌋
+
+private
+
+ -- The following unit test does not type-check (at the time of
+ -- writing) if _==_ is replaced by primCharEquality.
+
+ data P : (Char → Bool) → Set where
+ p : (c : Char) → P (c ==_)
+
+ unit-test : P ('x' ==_)
+ unit-test = p _
+
+------------------------------------------------------------------------
+-- Decidable equality
+
+decSetoid : DecSetoid _ _
+decSetoid = PropEq.decSetoid _≟_
diff --git a/src/Data/Cofin.agda b/src/Data/Cofin.agda
deleted file mode 100644
index 83e6d03..0000000
--- a/src/Data/Cofin.agda
+++ /dev/null
@@ -1,42 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- "Finite" sets indexed on coinductive "natural" numbers
-------------------------------------------------------------------------
-
-module Data.Cofin where
-
-open import Coinduction
-open import Data.Conat as Conat using (Coℕ; suc; ∞ℕ)
-open import Data.Nat.Base using (ℕ; zero; suc)
-open import Data.Fin using (Fin; zero; suc)
-
-------------------------------------------------------------------------
--- The type
-
--- Note that Cofin ∞ℕ is /not/ finite. Note also that this is not a
--- coinductive type, but it is indexed on a coinductive type.
-
-data Cofin : Coℕ → Set where
- zero : ∀ {n} → Cofin (suc n)
- suc : ∀ {n} (i : Cofin (♭ n)) → Cofin (suc n)
-
-------------------------------------------------------------------------
--- Some operations
-
-fromℕ : ℕ → Cofin ∞ℕ
-fromℕ zero = zero
-fromℕ (suc n) = suc (fromℕ n)
-
-toℕ : ∀ {n} → Cofin n → ℕ
-toℕ zero = zero
-toℕ (suc i) = suc (toℕ i)
-
-fromFin : ∀ {n} → Fin n → Cofin (Conat.fromℕ n)
-fromFin zero = zero
-fromFin (suc i) = suc (fromFin i)
-
-toFin : ∀ n → Cofin (Conat.fromℕ n) → Fin n
-toFin zero ()
-toFin (suc n) zero = zero
-toFin (suc n) (suc i) = suc (toFin n i)
diff --git a/src/Data/Colist.agda b/src/Data/Colist.agda
deleted file mode 100644
index c49503e..0000000
--- a/src/Data/Colist.agda
+++ /dev/null
@@ -1,501 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- Coinductive lists
-------------------------------------------------------------------------
-
-module Data.Colist where
-
-open import Category.Monad
-open import Coinduction
-open import Data.Bool.Base using (Bool; true; false)
-open import Data.BoundedVec.Inefficient as BVec
- using (BoundedVec; []; _∷_)
-open import Data.Conat using (Coℕ; zero; suc)
-open import Data.Empty using (⊥)
-open import Data.Maybe.Base using (Maybe; nothing; just; Is-just)
-open import Data.Nat.Base using (ℕ; zero; suc; _≥′_; ≤′-refl; ≤′-step)
-open import Data.Nat.Properties using (s≤′s)
-open import Data.List.Base using (List; []; _∷_)
-open import Data.List.NonEmpty using (List⁺; _∷_)
-open import Data.Product as Prod using (∃; _×_; _,_)
-open import Data.Sum as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Inverse as Inv using (_↔_; module Inverse)
-open import Level using (_⊔_)
-open import Relation.Binary
-import Relation.Binary.InducedPreorders as Ind
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
-open import Relation.Nullary
-open import Relation.Nullary.Negation
-
-module ¬¬Monad {p} where
- open RawMonad (¬¬-Monad {p}) public
-open ¬¬Monad -- we don't want the RawMonad content to be opened publicly
-
-------------------------------------------------------------------------
--- The type
-
-infixr 5 _∷_
-
-data Colist {a} (A : Set a) : Set a where
- [] : Colist A
- _∷_ : (x : A) (xs : ∞ (Colist A)) → Colist A
-
-{-# 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
- here : ∀ {x xs} (px : P x) → Any P (x ∷ xs)
- there : ∀ {x xs} (pxs : Any P (♭ xs)) → Any P (x ∷ xs)
-
-data All {a p} {A : Set a} (P : A → Set p) :
- Colist A → Set (a ⊔ p) where
- [] : All P []
- _∷_ : ∀ {x xs} (px : P x) (pxs : ∞ (All P (♭ xs))) → All P (x ∷ xs)
-
-------------------------------------------------------------------------
--- Some operations
-
-null : ∀ {a} {A : Set a} → Colist A → Bool
-null [] = true
-null (_ ∷ _) = false
-
-length : ∀ {a} {A : Set a} → Colist A → Coℕ
-length [] = zero
-length (x ∷ xs) = suc (♯ length (♭ xs))
-
-map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → Colist A → Colist B
-map f [] = []
-map f (x ∷ xs) = f x ∷ ♯ map f (♭ xs)
-
-fromList : ∀ {a} {A : Set a} → List A → Colist A
-fromList [] = []
-fromList (x ∷ xs) = x ∷ ♯ fromList xs
-
-take : ∀ {a} {A : Set a} (n : ℕ) → Colist A → BoundedVec A n
-take zero xs = []
-take (suc n) [] = []
-take (suc n) (x ∷ xs) = x ∷ take n (♭ xs)
-
-replicate : ∀ {a} {A : Set a} → Coℕ → A → Colist A
-replicate zero x = []
-replicate (suc n) x = x ∷ ♯ replicate (♭ n) x
-
-lookup : ∀ {a} {A : Set a} → ℕ → Colist A → Maybe A
-lookup n [] = nothing
-lookup zero (x ∷ xs) = just x
-lookup (suc n) (x ∷ xs) = lookup n (♭ xs)
-
-infixr 5 _++_
-
-_++_ : ∀ {a} {A : Set a} → Colist A → Colist A → Colist A
-[] ++ ys = ys
-(x ∷ xs) ++ ys = x ∷ ♯ (♭ xs ++ ys)
-
--- Interleaves the two colists (until the shorter one, if any, has
--- been exhausted).
-
-infixr 5 _⋎_
-
-_⋎_ : ∀ {a} {A : Set a} → Colist A → Colist A → Colist A
-[] ⋎ ys = ys
-(x ∷ xs) ⋎ ys = x ∷ ♯ (ys ⋎ ♭ xs)
-
-concat : ∀ {a} {A : Set a} → Colist (List⁺ A) → Colist A
-concat [] = []
-concat ((x ∷ []) ∷ xss) = x ∷ ♯ concat (♭ xss)
-concat ((x ∷ (y ∷ xs)) ∷ xss) = x ∷ ♯ concat ((y ∷ xs) ∷ xss)
-
-[_] : ∀ {a} {A : Set a} → A → Colist A
-[ x ] = x ∷ ♯ []
-
-------------------------------------------------------------------------
--- Any lemmas
-
--- Any lemma for map.
-
-Any-map : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {xs} →
- Any P (map f xs) ↔ Any (P ∘ f) xs
-Any-map {P = P} {f} {xs} = record
- { to = P.→-to-⟶ (to xs)
- ; from = P.→-to-⟶ (from xs)
- ; inverse-of = record
- { left-inverse-of = from∘to xs
- ; right-inverse-of = to∘from xs
- }
- }
- where
- to : ∀ xs → Any P (map f xs) → Any (P ∘ f) xs
- to [] ()
- to (x ∷ xs) (here px) = here px
- to (x ∷ xs) (there p) = there (to (♭ xs) p)
-
- from : ∀ xs → Any (P ∘ f) xs → Any P (map f xs)
- from [] ()
- from (x ∷ xs) (here px) = here px
- from (x ∷ xs) (there p) = there (from (♭ xs) p)
-
- from∘to : ∀ xs (p : Any P (map f xs)) → from xs (to xs p) ≡ p
- from∘to [] ()
- from∘to (x ∷ xs) (here px) = P.refl
- from∘to (x ∷ xs) (there p) = P.cong there (from∘to (♭ xs) p)
-
- to∘from : ∀ xs (p : Any (P ∘ f) xs) → to xs (from xs p) ≡ p
- to∘from [] ()
- to∘from (x ∷ xs) (here px) = P.refl
- to∘from (x ∷ xs) (there p) = P.cong there (to∘from (♭ xs) p)
-
--- Any lemma for _⋎_. This lemma implies that every member of xs or ys
--- is a member of xs ⋎ ys, and vice versa.
-
-Any-⋎ : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} →
- Any P (xs ⋎ ys) ↔ (Any P xs ⊎ Any P ys)
-Any-⋎ {P = P} = λ xs → record
- { to = P.→-to-⟶ (to xs)
- ; from = P.→-to-⟶ (from xs)
- ; inverse-of = record
- { left-inverse-of = from∘to xs
- ; right-inverse-of = to∘from xs
- }
- }
- where
- to : ∀ xs {ys} → Any P (xs ⋎ ys) → Any P xs ⊎ Any P ys
- to [] p = inj₂ p
- to (x ∷ xs) (here px) = inj₁ (here px)
- to (x ∷ xs) (there p) = [ inj₂ , inj₁ ∘ there ]′ (to _ p)
-
- mutual
-
- from-left : ∀ {xs ys} → Any P xs → Any P (xs ⋎ ys)
- from-left (here px) = here px
- from-left {ys = ys} (there p) = there (from-right ys p)
-
- from-right : ∀ xs {ys} → Any P ys → Any P (xs ⋎ ys)
- from-right [] p = p
- from-right (x ∷ xs) p = there (from-left p)
-
- from : ∀ xs {ys} → Any P xs ⊎ Any P ys → Any P (xs ⋎ ys)
- from xs = Sum.[ from-left , from-right xs ]
-
- from∘to : ∀ xs {ys} (p : Any P (xs ⋎ ys)) → from xs (to xs p) ≡ p
- from∘to [] p = P.refl
- from∘to (x ∷ xs) (here px) = P.refl
- from∘to (x ∷ xs) {ys = ys} (there p) with to ys p | from∘to ys p
- from∘to (x ∷ xs) {ys = ys} (there .(from-left q)) | inj₁ q | P.refl = P.refl
- from∘to (x ∷ xs) {ys = ys} (there .(from-right ys q)) | inj₂ q | P.refl = P.refl
-
- mutual
-
- to∘from-left : ∀ {xs ys} (p : Any P xs) →
- to xs {ys = ys} (from-left p) ≡ inj₁ p
- to∘from-left (here px) = P.refl
- to∘from-left {ys = ys} (there p)
- rewrite to∘from-right ys p = P.refl
-
- to∘from-right : ∀ xs {ys} (p : Any P ys) →
- to xs (from-right xs p) ≡ inj₂ p
- to∘from-right [] p = P.refl
- to∘from-right (x ∷ xs) {ys = ys} p
- rewrite to∘from-left {xs = ys} {ys = ♭ xs} p = P.refl
-
- to∘from : ∀ xs {ys} (p : Any P xs ⊎ Any P ys) → to xs (from xs p) ≡ p
- to∘from xs = Sum.[ to∘from-left , to∘from-right xs ]
-
-------------------------------------------------------------------------
--- Equality
-
--- xs ≈ ys means that xs and ys are equal.
-
-infix 4 _≈_
-
-data _≈_ {a} {A : Set a} : (xs ys : Colist A) → Set a where
- [] : [] ≈ []
- _∷_ : ∀ x {xs ys} (xs≈ : ∞ (♭ xs ≈ ♭ ys)) → x ∷ xs ≈ x ∷ ys
-
--- The equality relation forms a setoid.
-
-setoid : ∀ {ℓ} → Set ℓ → Setoid _ _
-setoid A = record
- { Carrier = Colist A
- ; _≈_ = _≈_
- ; isEquivalence = record
- { refl = refl
- ; sym = sym
- ; trans = trans
- }
- }
- where
- refl : Reflexive _≈_
- refl {[]} = []
- refl {x ∷ xs} = x ∷ ♯ refl
-
- sym : Symmetric _≈_
- sym [] = []
- sym (x ∷ xs≈) = x ∷ ♯ sym (♭ xs≈)
-
- trans : Transitive _≈_
- trans [] [] = []
- trans (x ∷ xs≈) (.x ∷ ys≈) = x ∷ ♯ trans (♭ xs≈) (♭ ys≈)
-
-module ≈-Reasoning where
- import Relation.Binary.EqReasoning as EqR
- private
- open module R {a} {A : Set a} = EqR (setoid A) public
-
--- map preserves equality.
-
-map-cong : ∀ {a b} {A : Set a} {B : Set b}
- (f : A → B) → _≈_ =[ map f ]⇒ _≈_
-map-cong f [] = []
-map-cong f (x ∷ xs≈) = f x ∷ ♯ map-cong f (♭ xs≈)
-
--- Any respects pointwise implication (for the predicate) and equality
--- (for the colist).
-
-Any-resp :
- ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {xs ys} →
- (∀ {x} → P x → Q x) → xs ≈ ys → Any P xs → Any Q ys
-Any-resp f (x ∷ xs≈) (here px) = here (f px)
-Any-resp f (x ∷ xs≈) (there p) = there (Any-resp f (♭ xs≈) p)
-
--- Any maps pointwise isomorphic predicates and equal colists to
--- isomorphic types.
-
-Any-cong :
- ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {xs ys} →
- (∀ {x} → P x ↔ Q x) → xs ≈ ys → Any P xs ↔ Any Q ys
-Any-cong {A = A} P↔Q xs≈ys = record
- { to = P.→-to-⟶ (Any-resp (_⟨$⟩_ (Inverse.to P↔Q)) xs≈ys)
- ; from = P.→-to-⟶ (Any-resp (_⟨$⟩_ (Inverse.from P↔Q)) (sym xs≈ys))
- ; inverse-of = record
- { left-inverse-of = resp∘resp P↔Q xs≈ys (sym xs≈ys)
- ; right-inverse-of = resp∘resp (Inv.sym P↔Q) (sym xs≈ys) xs≈ys
- }
- }
- where
- open Setoid (setoid _) using (sym)
-
- resp∘resp : ∀ {p q} {P : A → Set p} {Q : A → Set q} {xs ys}
- (P↔Q : ∀ {x} → P x ↔ Q x)
- (xs≈ys : xs ≈ ys) (ys≈xs : ys ≈ xs) (p : Any P xs) →
- Any-resp (_⟨$⟩_ (Inverse.from P↔Q)) ys≈xs
- (Any-resp (_⟨$⟩_ (Inverse.to P↔Q)) xs≈ys p) ≡ p
- resp∘resp P↔Q (x ∷ xs≈) (.x ∷ ys≈) (here px) =
- P.cong here (Inverse.left-inverse-of P↔Q px)
- resp∘resp P↔Q (x ∷ xs≈) (.x ∷ ys≈) (there p) =
- P.cong there (resp∘resp P↔Q (♭ xs≈) (♭ ys≈) p)
-
-------------------------------------------------------------------------
--- Indices
-
--- Converts Any proofs to indices into the colist. The index can also
--- be seen as the size of the proof.
-
-index : ∀ {a p} {A : Set a} {P : A → Set p} {xs} → Any P xs → ℕ
-index (here px) = zero
-index (there p) = suc (index p)
-
--- The position returned by index is guaranteed to be within bounds.
-
-lookup-index : ∀ {a p} {A : Set a} {P : A → Set p} {xs} (p : Any P xs) →
- Is-just (lookup (index p) xs)
-lookup-index (here px) = just _
-lookup-index (there p) = lookup-index p
-
--- Any-resp preserves the index.
-
-index-Any-resp :
- ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q}
- {f : ∀ {x} → P x → Q x} {xs ys}
- (xs≈ys : xs ≈ ys) (p : Any P xs) →
- index (Any-resp f xs≈ys p) ≡ index p
-index-Any-resp (x ∷ xs≈) (here px) = P.refl
-index-Any-resp (x ∷ xs≈) (there p) =
- P.cong suc (index-Any-resp (♭ xs≈) p)
-
--- The left-to-right direction of Any-⋎ returns a proof whose size is
--- no larger than that of the input proof.
-
-index-Any-⋎ :
- ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} (p : Any P (xs ⋎ ys)) →
- index p ≥′ [ index , index ]′ (Inverse.to (Any-⋎ xs) ⟨$⟩ p)
-index-Any-⋎ [] p = ≤′-refl
-index-Any-⋎ (x ∷ xs) (here px) = ≤′-refl
-index-Any-⋎ (x ∷ xs) {ys = ys} (there p)
- with Inverse.to (Any-⋎ ys) ⟨$⟩ p | index-Any-⋎ ys p
-... | inj₁ q | q≤p = ≤′-step q≤p
-... | inj₂ q | q≤p = s≤′s q≤p
-
-------------------------------------------------------------------------
--- Memberships, subsets, prefixes
-
--- x ∈ xs means that x is a member of xs.
-
-infix 4 _∈_
-
-_∈_ : ∀ {a} → {A : Set a} → A → Colist A → Set a
-x ∈ xs = Any (_≡_ x) xs
-
--- xs ⊆ ys means that xs is a subset of ys.
-
-infix 4 _⊆_
-
-_⊆_ : ∀ {a} → {A : Set a} → Colist A → Colist A → Set a
-xs ⊆ ys = ∀ {x} → x ∈ xs → x ∈ ys
-
--- xs ⊑ ys means that xs is a prefix of ys.
-
-infix 4 _⊑_
-
-data _⊑_ {a} {A : Set a} : Colist A → Colist A → Set a where
- [] : ∀ {ys} → [] ⊑ ys
- _∷_ : ∀ x {xs ys} (p : ∞ (♭ xs ⊑ ♭ ys)) → x ∷ xs ⊑ x ∷ ys
-
--- Any can be expressed using _∈_ (and vice versa).
-
-Any-∈ : ∀ {a p} {A : Set a} {P : A → Set p} {xs} →
- Any P xs ↔ ∃ λ x → x ∈ xs × P x
-Any-∈ {P = P} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ (λ { (x , x∈xs , p) → from x∈xs p })
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = λ { (x , x∈xs , p) → to∘from x∈xs p }
- }
- }
- where
- to : ∀ {xs} → Any P xs → ∃ λ x → x ∈ xs × P x
- to (here p) = _ , here P.refl , p
- to (there p) = Prod.map id (Prod.map there id) (to p)
-
- from : ∀ {x xs} → x ∈ xs → P x → Any P xs
- from (here P.refl) p = here p
- from (there x∈xs) p = there (from x∈xs p)
-
- to∘from : ∀ {x xs} (x∈xs : x ∈ xs) (p : P x) →
- to (from x∈xs p) ≡ (x , x∈xs , p)
- to∘from (here P.refl) p = P.refl
- to∘from (there x∈xs) p =
- P.cong (Prod.map id (Prod.map there id)) (to∘from x∈xs p)
-
- from∘to : ∀ {xs} (p : Any P xs) →
- let (x , x∈xs , px) = to p in from x∈xs px ≡ p
- from∘to (here _) = P.refl
- from∘to (there p) = P.cong there (from∘to p)
-
--- Prefixes are subsets.
-
-⊑⇒⊆ : ∀ {a} → {A : Set a} → _⊑_ {A = A} ⇒ _⊆_
-⊑⇒⊆ [] ()
-⊑⇒⊆ (x ∷ xs⊑ys) (here ≡x) = here ≡x
-⊑⇒⊆ (_ ∷ xs⊑ys) (there x∈xs) = there (⊑⇒⊆ (♭ xs⊑ys) x∈xs)
-
--- The prefix relation forms a poset.
-
-⊑-Poset : ∀ {ℓ} → Set ℓ → Poset _ _ _
-⊑-Poset A = record
- { Carrier = Colist A
- ; _≈_ = _≈_
- ; _≤_ = _⊑_
- ; isPartialOrder = record
- { isPreorder = record
- { isEquivalence = Setoid.isEquivalence (setoid A)
- ; reflexive = reflexive
- ; trans = trans
- }
- ; antisym = antisym
- }
- }
- where
- reflexive : _≈_ ⇒ _⊑_
- reflexive [] = []
- reflexive (x ∷ xs≈) = x ∷ ♯ reflexive (♭ xs≈)
-
- trans : Transitive _⊑_
- trans [] _ = []
- trans (x ∷ xs≈) (.x ∷ ys≈) = x ∷ ♯ trans (♭ xs≈) (♭ ys≈)
-
- antisym : Antisymmetric _≈_ _⊑_
- antisym [] [] = []
- antisym (x ∷ p₁) (.x ∷ p₂) = x ∷ ♯ antisym (♭ p₁) (♭ p₂)
-
-module ⊑-Reasoning where
- import Relation.Binary.PartialOrderReasoning as POR
- private
- open module R {a} {A : Set a} = POR (⊑-Poset A)
- public renaming (_≤⟨_⟩_ to _⊑⟨_⟩_)
-
--- The subset relation forms a preorder.
-
-⊆-Preorder : ∀ {ℓ} → Set ℓ → Preorder _ _ _
-⊆-Preorder A =
- Ind.InducedPreorder₂ (setoid A) _∈_
- (λ xs≈ys → ⊑⇒⊆ (⊑P.reflexive xs≈ys))
- where module ⊑P = Poset (⊑-Poset A)
-
-module ⊆-Reasoning where
- import Relation.Binary.PreorderReasoning as PreR
- private
- open module R {a} {A : Set a} = PreR (⊆-Preorder A)
- public renaming (_∼⟨_⟩_ to _⊆⟨_⟩_)
-
- infix 1 _∈⟨_⟩_
-
- _∈⟨_⟩_ : ∀ {a} {A : Set a} (x : A) {xs ys} →
- x ∈ xs → xs IsRelatedTo ys → x ∈ ys
- x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
-
--- take returns a prefix.
-
-take-⊑ : ∀ {a} {A : Set a} n (xs : Colist A) →
- let toColist = fromList {a} ∘ BVec.toList in
- toColist (take n xs) ⊑ xs
-take-⊑ zero xs = []
-take-⊑ (suc n) [] = []
-take-⊑ (suc n) (x ∷ xs) = x ∷ ♯ take-⊑ n (♭ xs)
-
-------------------------------------------------------------------------
--- Finiteness and infiniteness
-
--- Finite xs means that xs has finite length.
-
-data Finite {a} {A : Set a} : Colist A → Set a where
- [] : Finite []
- _∷_ : ∀ x {xs} (fin : Finite (♭ xs)) → Finite (x ∷ xs)
-
--- Infinite xs means that xs has infinite length.
-
-data Infinite {a} {A : Set a} : Colist A → Set a where
- _∷_ : ∀ x {xs} (inf : ∞ (Infinite (♭ xs))) → Infinite (x ∷ xs)
-
--- Colists which are not finite are infinite.
-
-not-finite-is-infinite :
- ∀ {a} {A : Set a} (xs : Colist A) → ¬ Finite xs → Infinite xs
-not-finite-is-infinite [] hyp with hyp []
-... | ()
-not-finite-is-infinite (x ∷ xs) hyp =
- x ∷ ♯ not-finite-is-infinite (♭ xs) (hyp ∘ _∷_ x)
-
--- Colists are either finite or infinite (classically).
-
-finite-or-infinite :
- ∀ {a} {A : Set a} (xs : Colist A) → ¬ ¬ (Finite xs ⊎ Infinite xs)
-finite-or-infinite xs = helper <$> excluded-middle
- where
- helper : Dec (Finite xs) → Finite xs ⊎ Infinite xs
- helper (yes fin) = inj₁ fin
- helper (no ¬fin) = inj₂ $ not-finite-is-infinite xs ¬fin
-
--- Colists are not both finite and infinite.
-
-not-finite-and-infinite :
- ∀ {a} {A : Set a} {xs : Colist A} → Finite xs → Infinite xs → ⊥
-not-finite-and-infinite [] ()
-not-finite-and-infinite (x ∷ fin) (.x ∷ inf) =
- not-finite-and-infinite fin (♭ inf)
diff --git a/src/Data/Colist/Infinite-merge.agda b/src/Data/Colist/Infinite-merge.agda
deleted file mode 100644
index 977168e..0000000
--- a/src/Data/Colist/Infinite-merge.agda
+++ /dev/null
@@ -1,231 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- Infinite merge operation for coinductive lists
-------------------------------------------------------------------------
-
-module Data.Colist.Infinite-merge where
-
-open import Coinduction
-open import Data.Colist as Colist hiding (_⋎_)
-open import Data.Nat
-open import Data.Nat.Properties
-open import Data.Product as Prod
-open import Data.Sum
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Inverse as Inv using (_↔_; module Inverse)
-import Function.Related as Related
-open import Function.Related.TypeIsomorphisms
-open import Induction.Nat using (<′-well-founded)
-import Induction.WellFounded as WF
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
-open import Relation.Binary.Sum
-
--- Some code that is used to work around Agda's syntactic guardedness
--- checker.
-
-private
-
- infixr 5 _∷_ _⋎_
-
- data ColistP {a} (A : Set a) : Set a where
- [] : ColistP A
- _∷_ : A → ∞ (ColistP A) → ColistP A
- _⋎_ : ColistP A → ColistP A → ColistP A
-
- data ColistW {a} (A : Set a) : Set a where
- [] : ColistW A
- _∷_ : A → ColistP A → ColistW A
-
- program : ∀ {a} {A : Set a} → Colist A → ColistP A
- program [] = []
- program (x ∷ xs) = x ∷ ♯ program (♭ xs)
-
- mutual
-
- _⋎W_ : ∀ {a} {A : Set a} → ColistW A → ColistP A → ColistW A
- [] ⋎W ys = whnf ys
- (x ∷ xs) ⋎W ys = x ∷ (ys ⋎ xs)
-
- whnf : ∀ {a} {A : Set a} → ColistP A → ColistW A
- whnf [] = []
- whnf (x ∷ xs) = x ∷ ♭ xs
- whnf (xs ⋎ ys) = whnf xs ⋎W ys
-
- mutual
-
- ⟦_⟧P : ∀ {a} {A : Set a} → ColistP A → Colist A
- ⟦ xs ⟧P = ⟦ whnf xs ⟧W
-
- ⟦_⟧W : ∀ {a} {A : Set a} → ColistW A → Colist A
- ⟦ [] ⟧W = []
- ⟦ x ∷ xs ⟧W = x ∷ ♯ ⟦ xs ⟧P
-
- mutual
-
- ⋎-homP : ∀ {a} {A : Set a} (xs : ColistP A) {ys} →
- ⟦ xs ⋎ ys ⟧P ≈ ⟦ xs ⟧P Colist.⋎ ⟦ ys ⟧P
- ⋎-homP xs = ⋎-homW (whnf xs) _
-
- ⋎-homW : ∀ {a} {A : Set a} (xs : ColistW A) ys →
- ⟦ xs ⋎W ys ⟧W ≈ ⟦ xs ⟧W Colist.⋎ ⟦ ys ⟧P
- ⋎-homW (x ∷ xs) ys = x ∷ ♯ ⋎-homP ys
- ⋎-homW [] ys = begin ⟦ ys ⟧P ∎
- where open ≈-Reasoning
-
- ⟦program⟧P : ∀ {a} {A : Set a} (xs : Colist A) →
- ⟦ program xs ⟧P ≈ xs
- ⟦program⟧P [] = []
- ⟦program⟧P (x ∷ xs) = x ∷ ♯ ⟦program⟧P (♭ xs)
-
- Any-⋎P : ∀ {a p} {A : Set a} {P : A → Set p} xs {ys} →
- Any P ⟦ program xs ⋎ ys ⟧P ↔ (Any P xs ⊎ Any P ⟦ ys ⟧P)
- Any-⋎P {P = P} xs {ys} =
- Any P ⟦ program xs ⋎ ys ⟧P ↔⟨ Any-cong Inv.id (⋎-homP (program xs)) ⟩
- Any P (⟦ program xs ⟧P Colist.⋎ ⟦ ys ⟧P) ↔⟨ Any-⋎ _ ⟩
- (Any P ⟦ program xs ⟧P ⊎ Any P ⟦ ys ⟧P) ↔⟨ Any-cong Inv.id (⟦program⟧P _) ⊎-cong (_ ∎) ⟩
- (Any P xs ⊎ Any P ⟦ ys ⟧P) ∎
- where open Related.EquationalReasoning
-
- index-Any-⋎P :
- ∀ {a p} {A : Set a} {P : A → Set p} xs {ys}
- (p : Any P ⟦ program xs ⋎ ys ⟧P) →
- index p ≥′ [ index , index ]′ (Inverse.to (Any-⋎P xs) ⟨$⟩ p)
- index-Any-⋎P xs p
- with Any-resp id (⋎-homW (whnf (program xs)) _) p
- | index-Any-resp {f = id} (⋎-homW (whnf (program xs)) _) p
- index-Any-⋎P xs p | q | q≡p
- with Inverse.to (Any-⋎ ⟦ program xs ⟧P) ⟨$⟩ q
- | index-Any-⋎ ⟦ program xs ⟧P q
- index-Any-⋎P xs p | q | q≡p | inj₂ r | r≤q rewrite q≡p = r≤q
- index-Any-⋎P xs p | q | q≡p | inj₁ r | r≤q
- with Any-resp id (⟦program⟧P xs) r
- | index-Any-resp {f = id} (⟦program⟧P xs) r
- index-Any-⋎P xs p | q | q≡p | inj₁ r | r≤q | s | s≡r
- rewrite s≡r | q≡p = r≤q
-
--- Infinite variant of _⋎_.
-
-private
-
- merge′ : ∀ {a} {A : Set a} → Colist (A × Colist A) → ColistP A
- merge′ [] = []
- merge′ ((x , xs) ∷ xss) = x ∷ ♯ (program xs ⋎ merge′ (♭ xss))
-
-merge : ∀ {a} {A : Set a} → Colist (A × Colist A) → Colist A
-merge xss = ⟦ merge′ xss ⟧P
-
--- Any lemma for merge.
-
-Any-merge :
- ∀ {a p} {A : Set a} {P : A → Set p} xss →
- Any P (merge xss) ↔ Any (λ { (x , xs) → P x ⊎ Any P xs }) xss
-Any-merge {P = P} = λ xss → record
- { to = P.→-to-⟶ (proj₁ ∘ to xss)
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = proj₂ ∘ to xss
- ; right-inverse-of = λ p → begin
- proj₁ (to xss (from p)) ≡⟨ from-injective _ _ (proj₂ (to xss (from p))) ⟩
- p ∎
- }
- }
- where
- open P.≡-Reasoning
-
- -- The from function.
-
- Q = λ { (x , xs) → P x ⊎ Any P xs }
-
- from : ∀ {xss} → Any Q xss → Any P (merge xss)
- from (here (inj₁ p)) = here p
- from (here (inj₂ p)) = there (Inverse.from (Any-⋎P _) ⟨$⟩ inj₁ p)
- from (there {x = _ , xs} p) = there (Inverse.from (Any-⋎P xs) ⟨$⟩ inj₂ (from p))
-
- -- Some lemmas.
-
- drop-there :
- ∀ {a p} {A : Set a} {P : A → Set p} {x xs} {p q : Any P _} →
- _≡_ {A = Any P (x ∷ xs)} (there p) (there q) → p ≡ q
- drop-there P.refl = P.refl
-
- drop-inj₁ : ∀ {a b} {A : Set a} {B : Set b} {x y} →
- _≡_ {A = A ⊎ B} (inj₁ x) (inj₁ y) → x ≡ y
- drop-inj₁ P.refl = P.refl
-
- drop-inj₂ : ∀ {a b} {A : Set a} {B : Set b} {x y} →
- _≡_ {A = A ⊎ B} (inj₂ x) (inj₂ y) → x ≡ y
- drop-inj₂ P.refl = P.refl
-
- -- The from function is injective.
-
- from-injective : ∀ {xss} (p₁ p₂ : Any Q xss) →
- from p₁ ≡ from p₂ → p₁ ≡ p₂
- from-injective (here (inj₁ p)) (here (inj₁ .p)) P.refl = P.refl
- from-injective (here (inj₁ _)) (here (inj₂ _)) ()
- from-injective (here (inj₂ _)) (here (inj₁ _)) ()
- from-injective (here (inj₂ p₁)) (here (inj₂ p₂)) eq =
- P.cong (here ∘ inj₂) $
- drop-inj₁ $
- Inverse.injective (Inv.sym (Any-⋎P _)) {x = inj₁ p₁} {y = inj₁ p₂} $
- drop-there eq
- from-injective (here (inj₁ _)) (there _) ()
- from-injective (here (inj₂ p₁)) (there p₂) eq
- with Inverse.injective (Inv.sym (Any-⋎P _))
- {x = inj₁ p₁} {y = inj₂ (from p₂)}
- (drop-there eq)
- ... | ()
- from-injective (there _) (here (inj₁ _)) ()
- from-injective (there p₁) (here (inj₂ p₂)) eq
- with Inverse.injective (Inv.sym (Any-⋎P _))
- {x = inj₂ (from p₁)} {y = inj₁ p₂}
- (drop-there eq)
- ... | ()
- from-injective (there {x = _ , xs} p₁) (there p₂) eq =
- P.cong there $
- from-injective p₁ p₂ $
- drop-inj₂ $
- Inverse.injective (Inv.sym (Any-⋎P xs))
- {x = inj₂ (from p₁)} {y = inj₂ (from p₂)} $
- drop-there eq
-
- -- The to function (defined as a right inverse of from).
-
- Input = ∃ λ xss → Any P (merge xss)
-
- Pred : Input → Set _
- Pred (xss , p) = ∃ λ (q : Any Q xss) → from q ≡ p
-
- to : ∀ xss p → Pred (xss , p)
- to = λ xss p →
- WF.All.wfRec (WF.Inverse-image.well-founded size <′-well-founded) _
- Pred step (xss , p)
- where
- size : Input → ℕ
- size (_ , p) = index p
-
- step : ∀ p → WF.WfRec (_<′_ on size) Pred p → Pred p
- step ([] , ()) rec
- step ((x , xs) ∷ xss , here p) rec = here (inj₁ p) , P.refl
- step ((x , xs) ∷ xss , there p) rec
- with Inverse.to (Any-⋎P xs) ⟨$⟩ p
- | Inverse.left-inverse-of (Any-⋎P xs) p
- | index-Any-⋎P xs p
- step ((x , xs) ∷ xss , there .(Inverse.from (Any-⋎P xs) ⟨$⟩ inj₁ q)) rec | inj₁ q | P.refl | _ = here (inj₂ q) , P.refl
- step ((x , xs) ∷ xss , there .(Inverse.from (Any-⋎P xs) ⟨$⟩ inj₂ q)) rec | inj₂ q | P.refl | q≤p =
- Prod.map there
- (P.cong (there ∘ _⟨$⟩_ (Inverse.from (Any-⋎P xs)) ∘ inj₂))
- (rec (♭ xss , q) (s≤′s q≤p))
-
--- Every member of xss is a member of merge xss, and vice versa (with
--- equal multiplicities).
-
-∈-merge : ∀ {a} {A : Set a} {y : A} xss →
- y ∈ merge xss ↔ ∃₂ λ x xs → (x , xs) ∈ xss × (y ≡ x ⊎ y ∈ xs)
-∈-merge {y = y} xss =
- y ∈ merge xss ↔⟨ Any-merge _ ⟩
- Any (λ { (x , xs) → y ≡ x ⊎ y ∈ xs }) xss ↔⟨ Any-∈ ⟩
- (∃ λ { (x , xs) → (x , xs) ∈ xss × (y ≡ x ⊎ y ∈ xs) }) ↔⟨ Σ-assoc ⟩
- (∃₂ λ x xs → (x , xs) ∈ xss × (y ≡ x ⊎ y ∈ xs)) ∎
- where open Related.EquationalReasoning
diff --git a/src/Data/Conat.agda b/src/Data/Conat.agda
deleted file mode 100644
index c394e39..0000000
--- a/src/Data/Conat.agda
+++ /dev/null
@@ -1,64 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- Coinductive "natural" numbers
-------------------------------------------------------------------------
-
-module Data.Conat where
-
-open import Coinduction
-open import Data.Nat.Base using (ℕ; zero; suc)
-open import Relation.Binary
-
-------------------------------------------------------------------------
--- The type
-
-data Coℕ : Set where
- zero : Coℕ
- suc : (n : ∞ Coℕ) → Coℕ
-
-------------------------------------------------------------------------
--- Some operations
-
-fromℕ : ℕ → Coℕ
-fromℕ zero = zero
-fromℕ (suc n) = suc (♯ fromℕ n)
-
-∞ℕ : Coℕ
-∞ℕ = suc (♯ ∞ℕ)
-
-infixl 6 _+_
-
-_+_ : Coℕ → Coℕ → Coℕ
-zero + n = n
-suc m + n = suc (♯ (♭ m + n))
-
-------------------------------------------------------------------------
--- Equality
-
-data _≈_ : Coℕ → Coℕ → Set where
- zero : zero ≈ zero
- suc : ∀ {m n} (m≈n : ∞ (♭ m ≈ ♭ n)) → suc m ≈ suc n
-
-setoid : Setoid _ _
-setoid = record
- { Carrier = Coℕ
- ; _≈_ = _≈_
- ; isEquivalence = record
- { refl = refl
- ; sym = sym
- ; trans = trans
- }
- }
- where
- refl : Reflexive _≈_
- refl {zero} = zero
- refl {suc n} = suc (♯ refl)
-
- sym : Symmetric _≈_
- sym zero = zero
- sym (suc m≈n) = suc (♯ sym (♭ m≈n))
-
- trans : Transitive _≈_
- trans zero zero = zero
- trans (suc m≈n) (suc n≈k) = suc (♯ trans (♭ m≈n) (♭ n≈k))
diff --git a/src/Data/Container.agda b/src/Data/Container.agda
index b0897ed..d4f45de 100644
--- a/src/Data/Container.agda
+++ b/src/Data/Container.agda
@@ -6,19 +6,19 @@
module Data.Container where
-open import Data.M
+open import Codata.Musical.M hiding (map)
open import Data.Product as Prod hiding (map)
-open import Data.W
+open import Data.W hiding (map)
open import Function renaming (id to ⟨id⟩; _∘_ to _⟨∘⟩_)
open import Function.Equality using (_⟨$⟩_)
open import Function.Inverse using (_↔_; module Inverse)
import Function.Related as Related
open import Level
open import Relation.Binary
- using (Setoid; module Setoid; Preorder; module Preorder)
+ using (REL ; IsEquivalence; Setoid; module Setoid; Preorder; module Preorder)
open import Relation.Binary.PropositionalEquality as P
using (_≡_; _≗_; refl)
-open import Relation.Unary using (_⊆_)
+open import Relation.Unary using (Pred ; _⊆_)
------------------------------------------------------------------------
-- Containers
@@ -26,34 +26,22 @@ open import Relation.Unary using (_⊆_)
-- A container is a set of shapes, and for every shape a set of
-- positions.
-infix 5 _▷_
-
-record Container (ℓ : Level) : Set (suc ℓ) where
- constructor _▷_
- field
- Shape : Set ℓ
- Position : Shape → Set ℓ
-
+open import Data.Container.Core public
open Container public
--- The semantics ("extension") of a container.
-
-⟦_⟧ : ∀ {ℓ₁ ℓ₂} → Container ℓ₁ → Set ℓ₂ → Set (ℓ₁ ⊔ ℓ₂)
-⟦ C ⟧ X = Σ[ s ∈ Shape C ] (Position C s → X)
-
-- The least and greatest fixpoints of a container.
-μ : ∀ {ℓ} → Container ℓ → Set ℓ
-μ C = W (Shape C) (Position C)
+μ : ∀ {s p} → Container s p → Set (s ⊔ p)
+μ = W
-ν : ∀ {ℓ} → Container ℓ → Set ℓ
-ν C = M (Shape C) (Position C)
+ν : ∀ {s p} → Container s p → Set (s ⊔ p)
+ν = M
-- Equality, parametrised on an underlying relation.
-Eq : ∀ {c ℓ} {C : Container c} {X Y : Set c} →
- (X → Y → Set ℓ) → ⟦ C ⟧ X → ⟦ C ⟧ Y → Set (c ⊔ ℓ)
-Eq {C = C} _≈_ (s , f) (s′ , f′) =
+Eq : ∀ {s p x y e} (C : Container s p) {X : Set x} {Y : Set y} →
+ (REL X Y e) → ⟦ C ⟧ X → ⟦ C ⟧ Y → Set (s ⊔ p ⊔ e)
+Eq C _≈_ (s , f) (s′ , f′) =
Σ[ eq ∈ s ≡ s′ ] (∀ p → f p ≈ f′ (P.subst (Position C) eq p))
private
@@ -61,140 +49,133 @@ private
-- Note that, if propositional equality were extensional, then
-- Eq _≡_ and _≡_ would coincide.
- Eq⇒≡ : ∀ {c} {C : Container c} {X : Set c} {xs ys : ⟦ C ⟧ X} →
- P.Extensionality c c → Eq _≡_ xs ys → xs ≡ ys
- Eq⇒≡ {xs = s , f} {ys = .s , f′} ext (refl , f≈f′) =
- P.cong (_,_ s) (ext f≈f′)
-
-setoid : ∀ {ℓ} → Container ℓ → Setoid ℓ ℓ → Setoid ℓ ℓ
-setoid C X = record
- { Carrier = ⟦ C ⟧ X.Carrier
- ; _≈_ = _≈_
- ; isEquivalence = record
- { refl = (refl , λ _ → X.refl)
+ Eq⇒≡ : ∀ {s p x} {C : Container s p} {X : Set x} {xs ys : ⟦ C ⟧ X} →
+ P.Extensionality p x → Eq C _≡_ xs ys → xs ≡ ys
+ Eq⇒≡ ext (refl , f≈f′) = P.cong -,_ (ext f≈f′)
+
+
+module _ {s p x e} (C : Container s p) (X : Setoid x e) where
+
+ private
+ module X = Setoid X
+ _≈_ = Eq C X._≈_
+
+ isEquivalence : IsEquivalence _≈_
+ isEquivalence = record
+ { refl = refl , λ p → X.refl
; sym = sym
; trans = λ {_ _ zs} → trans zs
- }
- }
- where
- module X = Setoid X
+ } where
- _≈_ = Eq X._≈_
+ sym : ∀ {xs ys} → xs ≈ ys → ys ≈ xs
+ sym (refl , f) = (refl , X.sym ⟨∘⟩ f)
- sym : {xs ys : ⟦ C ⟧ X.Carrier} → xs ≈ ys → ys ≈ xs
- sym {_ , _} {._ , _} (refl , f) = (refl , X.sym ⟨∘⟩ f)
+ trans : ∀ {xs ys} zs → xs ≈ ys → ys ≈ zs → xs ≈ zs
+ trans _ (refl , f₁) (refl , f₂) = refl , λ p → X.trans (f₁ p) (f₂ p)
- trans : ∀ {xs ys : ⟦ C ⟧ X.Carrier} zs → xs ≈ ys → ys ≈ zs → xs ≈ zs
- trans {_ , _} {._ , _} (._ , _) (refl , f₁) (refl , f₂) =
- (refl , λ p → X.trans (f₁ p) (f₂ p))
+ setoid : Setoid (s ⊔ p ⊔ x) (s ⊔ p ⊔ e)
+ setoid = record
+ { Carrier = ⟦ C ⟧ X.Carrier
+ ; _≈_ = _≈_
+ ; isEquivalence = isEquivalence
+ }
------------------------------------------------------------------------
-- Functoriality
-- Containers are functors.
-map : ∀ {c ℓ} {C : Container c} {X Y : Set ℓ} → (X → Y) → ⟦ C ⟧ X → ⟦ C ⟧ Y
-map f = Prod.map ⟨id⟩ (λ g → f ⟨∘⟩ g)
+map : ∀ {s p x y} {C : Container s p} {X : Set x} {Y : Set y} →
+ (X → Y) → ⟦ C ⟧ X → ⟦ C ⟧ Y
+map f = Prod.map₂ (f ⟨∘⟩_)
module Map where
- identity : ∀ {c} {C : Container c} X →
- let module X = Setoid X in
- (xs : ⟦ C ⟧ X.Carrier) → Eq X._≈_ (map ⟨id⟩ xs) xs
+ identity :
+ ∀ {s p x e} {C : Container s p} (X : Setoid x e) →
+ let module X = Setoid X in (xs : ⟦ C ⟧ X.Carrier) → Eq C X._≈_ (map ⟨id⟩ xs) xs
identity {C = C} X xs = Setoid.refl (setoid C X)
- composition : ∀ {c} {C : Container c} {X Y : Set c} Z →
- let module Z = Setoid Z in
- (f : Y → Z.Carrier) (g : X → Y) (xs : ⟦ C ⟧ X) →
- Eq Z._≈_ (map f (map g xs)) (map (f ⟨∘⟩ g) xs)
+ composition :
+ ∀ {s p x y z e} {C : Container s p} {X : Set x} {Y : Set y} (Z : Setoid z e) →
+ let module Z = Setoid Z in
+ (f : Y → Z.Carrier) (g : X → Y) (xs : ⟦ C ⟧ X) →
+ Eq C Z._≈_ (map f (map g xs)) (map (f ⟨∘⟩ g) xs)
composition {C = C} Z f g xs = Setoid.refl (setoid C Z)
------------------------------------------------------------------------
-- Container morphisms
--- Representation of container morphisms.
-
-record _⇒_ {c} (C₁ C₂ : Container c) : Set c where
- field
- shape : Shape C₁ → Shape C₂
- position : ∀ {s} → Position C₂ (shape s) → Position C₁ s
-
-open _⇒_ public
-
--- Interpretation of _⇒_.
-
-⟪_⟫ : ∀ {c ℓ} {C₁ C₂ : Container c} →
- C₁ ⇒ C₂ → {X : Set ℓ} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X
-⟪ m ⟫ xs = (shape m (proj₁ xs) , proj₂ xs ⟨∘⟩ position m)
-
module Morphism where
-- Naturality.
- Natural : ∀ {c} {C₁ C₂ : Container c} →
- (∀ {X} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X) → Set (suc c)
- Natural {c} {C₁} m =
- ∀ {X} (Y : Setoid c c) → let module Y = Setoid Y in
+ Natural : ∀ {s₁ s₂ p₁ p₂} x e {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂} →
+ (∀ {X : Set x} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X) →
+ Set (s₁ ⊔ s₂ ⊔ p₁ ⊔ p₂ ⊔ suc (x ⊔ e))
+ Natural x e {C₁ = C₁} {C₂} m =
+ ∀ {X : Set x} (Y : Setoid x e) → let module Y = Setoid Y in
(f : X → Y.Carrier) (xs : ⟦ C₁ ⟧ X) →
- Eq Y._≈_ (m $ map f xs) (map f $ m xs)
+ Eq C₂ Y._≈_ (m $ map f xs) (map f $ m xs)
-- Natural transformations.
- NT : ∀ {c} (C₁ C₂ : Container c) → Set (suc c)
- NT C₁ C₂ = ∃ λ (m : ∀ {X} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X) → Natural m
+ NT : ∀ {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) x e →
+ Set (s₁ ⊔ s₂ ⊔ p₁ ⊔ p₂ ⊔ suc (x ⊔ e))
+ NT C₁ C₂ x e = ∃ λ (m : ∀ {X : Set x} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X) → Natural x e m
-- Container morphisms are natural.
- natural : ∀ {c} {C₁ C₂ : Container c}
- (m : C₁ ⇒ C₂) → Natural ⟪ m ⟫
- natural {C₂ = C₂} m Y f xs = Setoid.refl (setoid C₂ Y)
+ natural : ∀ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂}
+ (m : C₁ ⇒ C₂) x e → Natural x e ⟪ m ⟫
+ natural m x e Y f xs = Setoid.refl (setoid _ Y)
-- In fact, all natural functions of the right type are container
-- morphisms.
- complete : ∀ {c} {C₁ C₂ : Container c} →
- (nt : NT C₁ C₂) →
- ∃ λ m → (X : Setoid c c) →
- let module X = Setoid X in
- (xs : ⟦ C₁ ⟧ X.Carrier) →
- Eq X._≈_ (proj₁ nt xs) (⟪ m ⟫ xs)
- complete (nt , nat) =
- (m , λ X xs → nat X (proj₂ xs) (proj₁ xs , ⟨id⟩))
- where
- m = record { shape = λ s → proj₁ (nt (s , ⟨id⟩))
- ; position = λ {s} → proj₂ (nt (s , ⟨id⟩))
- }
+ complete : ∀ {s₁ s₂ p₁ p₂ e} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂}
+ (nt : NT C₁ C₂ p₁ e) →
+ ∃ λ m → (X : Setoid p₁ e) → let module X = Setoid X in
+ ∀ xs → Eq C₂ X._≈_ (proj₁ nt xs) (⟪ m ⟫ xs)
+ complete {p₁ = p₁} {C₁ = C₁} {C₂} (nt , nat) =
+ (m , λ X xs → nat X (proj₂ xs) (proj₁ xs , ⟨id⟩)) where
+
+ m : C₁ ⇒ C₂
+ m .shape = λ s → proj₁ (nt (s , ⟨id⟩))
+ m .position = proj₂ (nt (_ , ⟨id⟩))
+
+
+ -- Combinators which commute with ⟪_⟫.
-- Identity.
- id : ∀ {c} (C : Container c) → C ⇒ C
- id _ = record {shape = ⟨id⟩; position = ⟨id⟩}
+ module _ {s p} (C : Container s p) where
- -- Composition.
+ id : C ⇒ C
+ id = ⟨id⟩ ▷ ⟨id⟩
- infixr 9 _∘_
- _∘_ : ∀ {c} {C₁ C₂ C₃ : Container c} → C₂ ⇒ C₃ → C₁ ⇒ C₂ → C₁ ⇒ C₃
- f ∘ g = record
- { shape = shape f ⟨∘⟩ shape g
- ; position = position g ⟨∘⟩ position f
- }
+ id-correct : ∀ {x} {X : Set x} → ⟪ id ⟫ {X} ≗ ⟨id⟩
+ id-correct x = refl
+
+ -- Composition.
- -- Identity and composition commute with ⟪_⟫.
+ module _ {s₁ s₂ s₃ p₁ p₂ p₃}
+ {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂} {C₃ : Container s₃ p₃} where
- id-correct : ∀ {c} {C : Container c} {X : Set c} →
- ⟪ id C ⟫ {X} ≗ ⟨id⟩
- id-correct xs = refl
+ infixr 9 _∘_
+ _∘_ : C₂ ⇒ C₃ → C₁ ⇒ C₂ → C₁ ⇒ C₃
+ (f ∘ g) .shape = shape f ⟨∘⟩ shape g
+ (f ∘ g) .position = position g ⟨∘⟩ position f
- ∘-correct : ∀ {c} {C₁ C₂ C₃ : Container c}
- (f : C₂ ⇒ C₃) (g : C₁ ⇒ C₂) {X : Set c} →
- ⟪ f ∘ g ⟫ {X} ≗ (⟪ f ⟫ ⟨∘⟩ ⟪ g ⟫)
- ∘-correct f g xs = refl
+ ∘-correct : ∀ f g {x} {X : Set x} → ⟪ f ∘ g ⟫ {X} ≗ (⟪ f ⟫ ⟨∘⟩ ⟪ g ⟫)
+ ∘-correct f g xs = refl
------------------------------------------------------------------------
-- Linear container morphisms
-record _⊸_ {c} (C₁ C₂ : Container c) : Set c where
+record _⊸_ {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂)
+ : Set (s₁ ⊔ s₂ ⊔ p₁ ⊔ p₂) where
field
shape⊸ : Shape C₁ → Shape C₂
position⊸ : ∀ {s} → Position C₂ (shape⊸ s) ↔ Position C₁ s
@@ -213,33 +194,24 @@ open _⊸_ public using (shape⊸; position⊸; ⟪_⟫⊸)
------------------------------------------------------------------------
-- All and any
--- All.
+module _ {s p x} {C : Container s p} {X : Set x} where
-□ : ∀ {c} {C : Container c} {X : Set c} →
- (X → Set c) → (⟦ C ⟧ X → Set c)
-□ P (s , f) = ∀ p → P (f p)
+-- All.
-□-map : ∀ {c} {C : Container c} {X : Set c} {P Q : X → Set c} →
- P ⊆ Q → □ {C = C} P ⊆ □ Q
-□-map P⊆Q = _⟨∘⟩_ P⊆Q
+ □-map : ∀ {ℓ ℓ′} {P : Pred X ℓ} {Q : Pred X ℓ′} → P ⊆ Q → □ {C = C} P ⊆ □ Q
+ □-map P⊆Q = _⟨∘⟩_ P⊆Q
-- Any.
-◇ : ∀ {c} {C : Container c} {X : Set c} →
- (X → Set c) → (⟦ C ⟧ X → Set c)
-◇ P (s , f) = ∃ λ p → P (f p)
-
-◇-map : ∀ {c} {C : Container c} {X : Set c} {P Q : X → Set c} →
- P ⊆ Q → ◇ {C = C} P ⊆ ◇ Q
-◇-map P⊆Q = Prod.map ⟨id⟩ P⊆Q
+ ◇-map : ∀ {ℓ ℓ′} {P : Pred X ℓ} {Q : Pred X ℓ′} → P ⊆ Q → ◇ {C = C} P ⊆ ◇ Q
+ ◇-map P⊆Q = Prod.map ⟨id⟩ P⊆Q
-- Membership.
-infix 4 _∈_
+ infix 4 _∈_
-_∈_ : ∀ {c} {C : Container c} {X : Set c} →
- X → ⟦ C ⟧ X → Set c
-x ∈ xs = ◇ (_≡_ x) xs
+ _∈_ : X → ⟦ C ⟧ X → Set (p ⊔ x)
+ x ∈ xs = ◇ (_≡_ x) xs
-- Bag and set equality and related preorders. Two containers xs and
-- ys are equal when viewed as sets if, whenever x ∈ xs, we also have
@@ -256,14 +228,16 @@ open Related public
; bijection to bag
)
-[_]-Order : ∀ {ℓ} → Kind → Container ℓ → Set ℓ → Preorder ℓ ℓ ℓ
+[_]-Order : ∀ {s p ℓ} → Kind → Container s p → Set ℓ →
+ Preorder (s ⊔ p ⊔ ℓ) (s ⊔ p ⊔ ℓ) (p ⊔ ℓ)
[ k ]-Order C X = Related.InducedPreorder₂ k (_∈_ {C = C} {X = X})
-[_]-Equality : ∀ {ℓ} → Symmetric-kind → Container ℓ → Set ℓ → Setoid ℓ ℓ
+[_]-Equality : ∀ {s p ℓ} → Symmetric-kind → Container s p → Set ℓ →
+ Setoid (s ⊔ p ⊔ ℓ) (p ⊔ ℓ)
[ k ]-Equality C X = Related.InducedEquivalence₂ k (_∈_ {C = C} {X = X})
infix 4 _∼[_]_
-_∼[_]_ : ∀ {c} {C : Container c} {X : Set c} →
- ⟦ C ⟧ X → Kind → ⟦ C ⟧ X → Set c
+_∼[_]_ : ∀ {s p x} {C : Container s p} {X : Set x} →
+ ⟦ C ⟧ X → Kind → ⟦ C ⟧ X → Set (p ⊔ x)
_∼[_]_ {C = C} {X} xs k ys = Preorder._∼_ ([ k ]-Order C X) xs ys
diff --git a/src/Data/Container/Any.agda b/src/Data/Container/Any.agda
index 4f52fe2..5624856 100644
--- a/src/Data/Container/Any.agda
+++ b/src/Data/Container/Any.agda
@@ -6,256 +6,244 @@
module Data.Container.Any where
+open import Level
open import Algebra
open import Data.Container as C
open import Data.Container.Combinator
using (module Composition) renaming (_∘_ to _⟨∘⟩_)
open import Data.Product as Prod hiding (swap)
-open import Data.Sum
+open import Data.Product.Relation.Pointwise.NonDependent
+import Data.Product.Relation.Pointwise.Dependent as Σ
+open import Data.Sum using (_⊎_; inj₁; inj₂; [_,_])
open import Function
open import Function.Equality using (_⟨$⟩_)
-open import Function.Inverse as Inv using (_↔_; module Inverse)
-open import Function.Related as Related using (Related)
+open import Function.Equivalence using (equivalence)
+open import Function.Inverse as Inv using (_↔_; inverse; module Inverse)
+open import Function.Related as Related using (Related; SK-sym)
open import Function.Related.TypeIsomorphisms
+open import Relation.Unary using (Pred ; _∪_ ; _∩_)
+open import Relation.Binary using (REL)
import Relation.Binary.HeterogeneousEquality as H
-open import Relation.Binary.Product.Pointwise
open import Relation.Binary.PropositionalEquality as P
using (_≡_; _≗_; refl)
-import Relation.Binary.Sigma.Pointwise as Σ
open Related.EquationalReasoning
private
- module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
+ module ×⊎ {k ℓ} = CommutativeSemiring (×-⊎-commutativeSemiring k ℓ)
+
+module _ {s p} (C : Container s p) {x} {X : Set x} {ℓ} {P : Pred X ℓ} where
-- ◇ can be expressed using _∈_.
-↔∈ : ∀ {c} (C : Container c) {X : Set c}
- {P : X → Set c} {xs : ⟦ C ⟧ X} →
- ◇ P xs ↔ (∃ λ x → x ∈ xs × P x)
-↔∈ _ {P = P} {xs} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = to∘from
- }
- }
- where
- to : ◇ P xs → ∃ λ x → x ∈ xs × P x
- to (p , Px) = (proj₂ xs p , (p , refl) , Px)
-
- from : (∃ λ x → x ∈ xs × P x) → ◇ P xs
- from (.(proj₂ xs p) , (p , refl) , Px) = (p , Px)
-
- to∘from : to ∘ from ≗ id
- to∘from (.(proj₂ xs p) , (p , refl) , Px) = refl
+ ↔∈ : ∀ {xs : ⟦ C ⟧ X} → ◇ P xs ↔ (∃ λ x → x ∈ xs × P x)
+ ↔∈ {xs} = inverse to from (λ _ → refl) (to∘from)
+ where
+
+ to : ◇ P xs → ∃ λ x → x ∈ xs × P x
+ to (p , Px) = (proj₂ xs p , (p , refl) , Px)
+
+ from : (∃ λ x → x ∈ xs × P x) → ◇ P xs
+ from (.(proj₂ xs p) , (p , refl) , Px) = (p , Px)
+
+ to∘from : to ∘ from ≗ id
+ to∘from (.(proj₂ xs p) , (p , refl) , Px) = refl
+module _ {s p} {C : Container s p} {x} {X : Set x}
+ {ℓ₁ ℓ₂} {P₁ : Pred X ℓ₁} {P₂ : Pred X ℓ₂} where
-- ◇ is a congruence for bag and set equality and related preorders.
-cong : ∀ {k c} {C : Container c}
- {X : Set c} {P₁ P₂ : X → Set c} {xs₁ xs₂ : ⟦ C ⟧ X} →
- (∀ x → Related k (P₁ x) (P₂ x)) → xs₁ ∼[ k ] xs₂ →
- Related k (◇ P₁ xs₁) (◇ P₂ xs₂)
-cong {C = C} {P₁ = P₁} {P₂} {xs₁} {xs₂} P₁↔P₂ xs₁≈xs₂ =
- ◇ P₁ xs₁ ↔⟨ ↔∈ C ⟩
- (∃ λ x → x ∈ xs₁ × P₁ x) ∼⟨ Σ.cong Inv.id (xs₁≈xs₂ ×-cong P₁↔P₂ _) ⟩
- (∃ λ x → x ∈ xs₂ × P₂ x) ↔⟨ sym (↔∈ C) ⟩
- ◇ P₂ xs₂ ∎
+ cong : ∀ {k} {xs₁ xs₂ : ⟦ C ⟧ X} →
+ (∀ x → Related k (P₁ x) (P₂ x)) → xs₁ ∼[ k ] xs₂ →
+ Related k (◇ P₁ xs₁) (◇ P₂ xs₂)
+ cong {k} {xs₁} {xs₂} P₁↔P₂ xs₁≈xs₂ =
+ ◇ P₁ xs₁ ↔⟨ ↔∈ C ⟩
+ (∃ λ x → x ∈ xs₁ × P₁ x) ∼⟨ Σ.cong Inv.id (xs₁≈xs₂ ×-cong P₁↔P₂ _) ⟩
+ (∃ λ x → x ∈ xs₂ × P₂ x) ↔⟨ SK-sym (↔∈ C) ⟩
+ ◇ P₂ xs₂ ∎
-- Nested occurrences of ◇ can sometimes be swapped.
-swap : ∀ {c} {C₁ C₂ : Container c} {X Y : Set c} {P : X → Y → Set c}
- {xs : ⟦ C₁ ⟧ X} {ys : ⟦ C₂ ⟧ Y} →
- let ◈ : ∀ {C : Container c} {X} → ⟦ C ⟧ X → (X → Set c) → Set c
- ◈ = λ {_} {_} → flip ◇ in
- ◈ xs (◈ ys ∘ P) ↔ ◈ ys (◈ xs ∘ flip P)
-swap {c} {C₁} {C₂} {P = P} {xs} {ys} =
- ◇ (λ x → ◇ (P x) ys) xs ↔⟨ ↔∈ C₁ ⟩
- (∃ λ x → x ∈ xs × ◇ (P x) ys) ↔⟨ Σ.cong Inv.id (λ {x} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ ↔∈ C₂ {P = P x}) ⟩
- (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ↔⟨ Σ.cong Inv.id (λ {x} → ∃∃↔∃∃ {A = x ∈ xs} (λ _ y → y ∈ ys × P x y)) ⟩
- (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ↔⟨ ∃∃↔∃∃ (λ x y → x ∈ xs × y ∈ ys × P x y) ⟩
- (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
- (x ∈ xs × y ∈ ys × P x y) ↔⟨ sym $ ×⊎.*-assoc _ _ _ ⟩
- ((x ∈ xs × y ∈ ys) × P x y) ↔⟨ ×⊎.*-comm _ _ ⟨ ×⊎.*-cong {ℓ = c} ⟩ Inv.id ⟩
- ((y ∈ ys × x ∈ xs) × P x y) ↔⟨ ×⊎.*-assoc _ _ _ ⟩
- (y ∈ ys × x ∈ xs × P x y) ∎)) ⟩
- (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → ∃∃↔∃∃ {B = y ∈ ys} (λ x _ → x ∈ xs × P x y)) ⟩
- (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → Inv.id ⟨ ×⊎.*-cong {ℓ = c} ⟩ sym (↔∈ C₁ {P = flip P y})) ⟩
- (∃ λ y → y ∈ ys × ◇ (flip P y) xs) ↔⟨ sym (↔∈ C₂) ⟩
- ◇ (λ y → ◇ (flip P y) xs) ys ∎
+module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂}
+ {x y} {X : Set x} {Y : Set y} {r} {P : REL X Y r} where
+
+ swap : {xs : ⟦ C₁ ⟧ X} {ys : ⟦ C₂ ⟧ Y} →
+ let ◈ : ∀ {s p} {C : Container s p} {x} {X : Set x} {ℓ} → ⟦ C ⟧ X → Pred X ℓ → Set (p ⊔ ℓ)
+ ◈ = λ {_} {_} → flip ◇ in
+ ◈ xs (◈ ys ∘ P) ↔ ◈ ys (◈ xs ∘ flip P)
+ swap {xs} {ys} =
+ ◇ (λ x → ◇ (P x) ys) xs ↔⟨ ↔∈ C₁ ⟩
+ (∃ λ x → x ∈ xs × ◇ (P x) ys) ↔⟨ Σ.cong Inv.id $ Σ.cong Inv.id $ ↔∈ C₂ ⟩
+ (∃ λ x → x ∈ xs × ∃ λ y → y ∈ ys × P x y) ↔⟨ Σ.cong Inv.id (λ {x} → ∃∃↔∃∃ (λ _ y → y ∈ ys × P x y)) ⟩
+ (∃₂ λ x y → x ∈ xs × y ∈ ys × P x y) ↔⟨ ∃∃↔∃∃ (λ x y → x ∈ xs × y ∈ ys × P x y) ⟩
+ (∃₂ λ y x → x ∈ xs × y ∈ ys × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → Σ.cong Inv.id (λ {x} →
+ (x ∈ xs × y ∈ ys × P x y) ↔⟨ SK-sym Σ-assoc ⟩
+ ((x ∈ xs × y ∈ ys) × P x y) ↔⟨ Σ.cong (×-comm _ _) Inv.id ⟩
+ ((y ∈ ys × x ∈ xs) × P x y) ↔⟨ Σ-assoc ⟩
+ (y ∈ ys × x ∈ xs × P x y) ∎)) ⟩
+ (∃₂ λ y x → y ∈ ys × x ∈ xs × P x y) ↔⟨ Σ.cong Inv.id (λ {y} → ∃∃↔∃∃ {B = y ∈ ys} (λ x _ → x ∈ xs × P x y)) ⟩
+ (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ↔⟨ Σ.cong Inv.id (Σ.cong Inv.id (SK-sym (↔∈ C₁))) ⟩
+ (∃ λ y → y ∈ ys × ◇ (flip P y) xs) ↔⟨ SK-sym (↔∈ C₂) ⟩
+ ◇ (λ y → ◇ (flip P y) xs) ys ∎
-- Nested occurrences of ◇ can sometimes be flattened.
-flatten : ∀ {c} {C₁ C₂ : Container c} {X}
- P (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) →
- ◇ (◇ P) xss ↔
- ◇ P (Inverse.from (Composition.correct C₁ C₂) ⟨$⟩ xss)
-flatten {C₁ = C₁} {C₂} {X} P xss = record
- { to = P.→-to-⟶ t
- ; from = P.→-to-⟶ f
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
- where
- open Inverse
-
- t : ◇ (◇ P) xss → ◇ P (from (Composition.correct C₁ C₂) ⟨$⟩ xss)
- t (p₁ , p₂ , p) = ((p₁ , p₂) , p)
-
- f : ◇ P (from (Composition.correct C₁ C₂) ⟨$⟩ xss) → ◇ (◇ P) xss
- f ((p₁ , p₂) , p) = (p₁ , p₂ , p)
+module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂}
+ {x} {X : Set x} {ℓ} (P : Pred X ℓ) where
+
+ flatten : ∀ (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) →
+ ◇ (◇ P) xss ↔
+ ◇ P (Inverse.from (Composition.correct C₁ C₂) ⟨$⟩ xss)
+ flatten xss = inverse t f (λ _ → refl) (λ _ → refl)
+ where
+ open Inverse
+
+ t : ◇ (◇ P) xss → ◇ P (from (Composition.correct C₁ C₂) ⟨$⟩ xss)
+ t (p₁ , p₂ , p) = ((p₁ , p₂) , p)
+
+ f : ◇ P (from (Composition.correct C₁ C₂) ⟨$⟩ xss) → ◇ (◇ P) xss
+ f ((p₁ , p₂) , p) = (p₁ , p₂ , p)
-- Sums commute with ◇ (for a fixed instance of a given container).
-◇⊎↔⊎◇ : ∀ {c} {C : Container c} {X : Set c} {xs : ⟦ C ⟧ X}
- {P Q : X → Set c} →
- ◇ (λ x → P x ⊎ Q x) xs ↔ (◇ P xs ⊎ ◇ Q xs)
-◇⊎↔⊎◇ {xs = xs} {P} {Q} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = [ (λ _ → refl) , (λ _ → refl) ]
- }
- }
- where
- to : ◇ (λ x → P x ⊎ Q x) xs → ◇ P xs ⊎ ◇ Q xs
- to (pos , inj₁ p) = inj₁ (pos , p)
- to (pos , inj₂ q) = inj₂ (pos , q)
-
- from : ◇ P xs ⊎ ◇ Q xs → ◇ (λ x → P x ⊎ Q x) xs
- from = [ Prod.map id inj₁ , Prod.map id inj₂ ]
-
- from∘to : from ∘ to ≗ id
- from∘to (pos , inj₁ p) = refl
- from∘to (pos , inj₂ q) = refl
+module _ {s p} {C : Container s p} {x} {X : Set x}
+ {ℓ ℓ′} {P : Pred X ℓ} {Q : Pred X ℓ′} where
+
+ ◇⊎↔⊎◇ : ∀ {xs : ⟦ C ⟧ X} → ◇ (P ∪ Q) xs ↔ (◇ P xs ⊎ ◇ Q xs)
+ ◇⊎↔⊎◇ {xs} = inverse to from from∘to to∘from
+ where
+ to : ◇ (λ x → P x ⊎ Q x) xs → ◇ P xs ⊎ ◇ Q xs
+ to (pos , inj₁ p) = inj₁ (pos , p)
+ to (pos , inj₂ q) = inj₂ (pos , q)
+
+ from : ◇ P xs ⊎ ◇ Q xs → ◇ (λ x → P x ⊎ Q x) xs
+ from = [ Prod.map id inj₁ , Prod.map id inj₂ ]
+
+ from∘to : from ∘ to ≗ id
+ from∘to (pos , inj₁ p) = refl
+ from∘to (pos , inj₂ q) = refl
+
+ to∘from : to ∘ from ≗ id
+ to∘from = [ (λ _ → refl) , (λ _ → refl) ]
-- Products "commute" with ◇.
-×◇↔◇◇× : ∀ {c} {C₁ C₂ : Container c}
- {X Y} {P : X → Set c} {Q : Y → Set c}
- {xs : ⟦ C₁ ⟧ X} {ys : ⟦ C₂ ⟧ Y} →
- ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs ↔ (◇ P xs × ◇ Q ys)
-×◇↔◇◇× {C₁ = C₁} {C₂} {P = P} {Q} {xs} {ys} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
- where
- to : ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs → ◇ P xs × ◇ Q ys
- to (p₁ , p₂ , p , q) = ((p₁ , p) , (p₂ , q))
-
- from : ◇ P xs × ◇ Q ys → ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs
- from ((p₁ , p) , (p₂ , q)) = (p₁ , p₂ , p , q)
+module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂}
+ {x y} {X : Set x} {Y : Set y} {ℓ ℓ′} {P : Pred X ℓ} {Q : Pred Y ℓ′} where
+
+ ×◇↔◇◇× : ∀ {xs : ⟦ C₁ ⟧ X} {ys : ⟦ C₂ ⟧ Y} →
+ ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs ↔ (◇ P xs × ◇ Q ys)
+ ×◇↔◇◇× {xs} {ys} = inverse to from (λ _ → refl) (λ _ → refl)
+ where
+ to : ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs → ◇ P xs × ◇ Q ys
+ to (p₁ , p₂ , p , q) = ((p₁ , p) , (p₂ , q))
+
+ from : ◇ P xs × ◇ Q ys → ◇ (λ x → ◇ (λ y → P x × Q y) ys) xs
+ from ((p₁ , p) , (p₂ , q)) = (p₁ , p₂ , p , q)
-- map can be absorbed by the predicate.
-map↔∘ : ∀ {c} (C : Container c) {X Y : Set c}
- (P : Y → Set c) {xs : ⟦ C ⟧ X} (f : X → Y) →
- ◇ P (C.map f xs) ↔ ◇ (P ∘ f) xs
-map↔∘ _ _ _ = Inv.id
+module _ {s p} (C : Container s p) {x y} {X : Set x} {Y : Set y}
+ {ℓ} (P : Pred Y ℓ) where
+
+ map↔∘ : ∀ {xs : ⟦ C ⟧ X} (f : X → Y) → ◇ P (C.map f xs) ↔ ◇ (P ∘ f) xs
+ map↔∘ f = Inv.id
-- Membership in a mapped container can be expressed without reference
-- to map.
-∈map↔∈×≡ : ∀ {c} (C : Container c) {X Y : Set c} {f : X → Y}
- {xs : ⟦ C ⟧ X} {y} →
- y ∈ C.map f xs ↔ (∃ λ x → x ∈ xs × y ≡ f x)
-∈map↔∈×≡ {c} C {f = f} {xs} {y} =
- y ∈ C.map f xs ↔⟨ map↔∘ C (_≡_ y) f ⟩
- ◇ (λ x → y ≡ f x) xs ↔⟨ ↔∈ C ⟩
- (∃ λ x → x ∈ xs × y ≡ f x) ∎
+module _ {s p} (C : Container s p) {x y} {X : Set x} {Y : Set y}
+ {ℓ} (P : Pred Y ℓ) where
+
+ ∈map↔∈×≡ : ∀ {f : X → Y} {xs : ⟦ C ⟧ X} {y} →
+ y ∈ C.map f xs ↔ (∃ λ x → x ∈ xs × y ≡ f x)
+ ∈map↔∈×≡ {f = f} {xs} {y} =
+ y ∈ C.map f xs ↔⟨ map↔∘ C (y ≡_) f ⟩
+ ◇ (λ x → y ≡ f x) xs ↔⟨ ↔∈ C ⟩
+ (∃ λ x → x ∈ xs × y ≡ f x) ∎
-- map is a congruence for bag and set equality and related preorders.
-map-cong : ∀ {k c} {C : Container c} {X Y : Set c}
- {f₁ f₂ : X → Y} {xs₁ xs₂ : ⟦ C ⟧ X} →
- f₁ ≗ f₂ → xs₁ ∼[ k ] xs₂ →
- C.map f₁ xs₁ ∼[ k ] C.map f₂ xs₂
-map-cong {c = c} {C} {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs₂ {x} =
- x ∈ C.map f₁ xs₁ ↔⟨ map↔∘ C (_≡_ x) f₁ ⟩
- ◇ (λ y → x ≡ f₁ y) xs₁ ∼⟨ cong {xs₁ = xs₁} {xs₂ = xs₂} (Related.↔⇒ ∘ helper) xs₁≈xs₂ ⟩
- ◇ (λ y → x ≡ f₂ y) xs₂ ↔⟨ sym (map↔∘ C (_≡_ x) f₂) ⟩
- x ∈ C.map f₂ xs₂ ∎
- where
- helper : ∀ y → (x ≡ f₁ y) ↔ (x ≡ f₂ y)
- helper y = record
- { to = P.→-to-⟶ (λ x≡f₁y → P.trans x≡f₁y ( f₁≗f₂ y))
- ; from = P.→-to-⟶ (λ x≡f₂y → P.trans x≡f₂y (P.sym $ f₁≗f₂ y))
- ; inverse-of = record
- { left-inverse-of = λ _ → P.proof-irrelevance _ _
- ; right-inverse-of = λ _ → P.proof-irrelevance _ _
+module _ {s p} (C : Container s p) {x y} {X : Set x} {Y : Set y}
+ {ℓ} (P : Pred Y ℓ) where
+
+ map-cong : ∀ {k} {f₁ f₂ : X → Y} {xs₁ xs₂ : ⟦ C ⟧ X} →
+ f₁ ≗ f₂ → xs₁ ∼[ k ] xs₂ →
+ C.map f₁ xs₁ ∼[ k ] C.map f₂ xs₂
+ map-cong {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs₂ {x} =
+ x ∈ C.map f₁ xs₁ ↔⟨ map↔∘ C (_≡_ x) f₁ ⟩
+ ◇ (λ y → x ≡ f₁ y) xs₁ ∼⟨ cong {xs₁ = xs₁} {xs₂ = xs₂} (Related.↔⇒ ∘ helper) xs₁≈xs₂ ⟩
+ ◇ (λ y → x ≡ f₂ y) xs₂ ↔⟨ SK-sym (map↔∘ C (_≡_ x) f₂) ⟩
+ x ∈ C.map f₂ xs₂ ∎
+ where
+ helper : ∀ y → (x ≡ f₁ y) ↔ (x ≡ f₂ y)
+ helper y = record
+ { to = P.→-to-⟶ (λ x≡f₁y → P.trans x≡f₁y ( f₁≗f₂ y))
+ ; from = P.→-to-⟶ (λ x≡f₂y → P.trans x≡f₂y (P.sym $ f₁≗f₂ y))
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.≡-irrelevance _ _
+ ; right-inverse-of = λ _ → P.≡-irrelevance _ _
+ }
}
- }
-- Uses of linear morphisms can be removed.
-remove-linear :
- ∀ {c} {C₁ C₂ : Container c} {X} {xs : ⟦ C₁ ⟧ X}
- (P : X → Set c) (m : C₁ ⊸ C₂) →
- ◇ P (⟪ m ⟫⊸ xs) ↔ ◇ P xs
-remove-linear {xs = xs} P m = record
- { to = P.→-to-⟶ t
- ; from = P.→-to-⟶ f
- ; inverse-of = record
- { left-inverse-of = f∘t
- ; right-inverse-of = t∘f
- }
- }
- where
- open Inverse
-
- t : ◇ P (⟪ m ⟫⊸ xs) → ◇ P xs
- t = Prod.map (_⟨$⟩_ (to (position⊸ m))) id
-
- f : ◇ P xs → ◇ P (⟪ m ⟫⊸ xs)
- f = Prod.map (_⟨$⟩_ (from (position⊸ m)))
- (P.subst (P ∘ proj₂ xs)
- (P.sym $ right-inverse-of (position⊸ m) _))
-
- f∘t : f ∘ t ≗ id
- f∘t (p₂ , p) = H.≅-to-≡ $
- H.cong₂ _,_ (H.≡-to-≅ $ left-inverse-of (position⊸ m) p₂)
- (H.≡-subst-removable
- (P ∘ proj₂ xs)
- (P.sym (right-inverse-of (position⊸ m)
- (to (position⊸ m) ⟨$⟩ p₂)))
- p)
-
- t∘f : t ∘ f ≗ id
- t∘f (p₁ , p) = H.≅-to-≡ $
- H.cong₂ _,_ (H.≡-to-≅ $ right-inverse-of (position⊸ m) p₁)
- (H.≡-subst-removable
- (P ∘ proj₂ xs)
- (P.sym (right-inverse-of (position⊸ m) p₁))
- p)
+module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂}
+ {x} {X : Set x} {ℓ} (P : Pred X ℓ) where
+
+ remove-linear : ∀ {xs : ⟦ C₁ ⟧ X} (m : C₁ ⊸ C₂) → ◇ P (⟪ m ⟫⊸ xs) ↔ ◇ P xs
+ remove-linear {xs} m = inverse t f f∘t t∘f
+ where
+ open Inverse
+
+ t : ◇ P (⟪ m ⟫⊸ xs) → ◇ P xs
+ t = Prod.map (to (position⊸ m) ⟨$⟩_) id
+
+ f : ◇ P xs → ◇ P (⟪ m ⟫⊸ xs)
+ f = Prod.map (from (position⊸ m) ⟨$⟩_)
+ (P.subst (P ∘ proj₂ xs)
+ (P.sym $ right-inverse-of (position⊸ m) _))
+
+ f∘t : f ∘ t ≗ id
+ f∘t (p₂ , p) = H.≅-to-≡ $
+ H.cong₂ _,_ (H.≡-to-≅ $ left-inverse-of (position⊸ m) p₂)
+ (H.≡-subst-removable
+ (P ∘ proj₂ xs)
+ (P.sym (right-inverse-of (position⊸ m)
+ (to (position⊸ m) ⟨$⟩ p₂)))
+ p)
+
+ t∘f : t ∘ f ≗ id
+ t∘f (p₁ , p) = H.≅-to-≡ $
+ H.cong₂ _,_ (H.≡-to-≅ $ right-inverse-of (position⊸ m) p₁)
+ (H.≡-subst-removable
+ (P ∘ proj₂ xs)
+ (P.sym (right-inverse-of (position⊸ m) p₁))
+ p)
-- Linear endomorphisms are identity functions if bag equality is
-- used.
-linear-identity :
- ∀ {c} {C : Container c} {X} {xs : ⟦ C ⟧ X} (m : C ⊸ C) →
- ⟪ m ⟫⊸ xs ∼[ bag ] xs
-linear-identity {xs = xs} m {x} =
- x ∈ ⟪ m ⟫⊸ xs ↔⟨ remove-linear (_≡_ x) m ⟩
- x ∈ xs ∎
+module _ {s p} {C : Container s p} {x} {X : Set x} where
+
+ linear-identity : ∀ {xs : ⟦ C ⟧ X} (m : C ⊸ C) → ⟪ m ⟫⊸ xs ∼[ bag ] xs
+ linear-identity {xs} m {x} =
+ x ∈ ⟪ m ⟫⊸ xs ↔⟨ remove-linear (_≡_ x) m ⟩
+ x ∈ xs ∎
-- If join can be expressed using a linear morphism (in a certain
-- way), then it can be absorbed by the predicate.
-join↔◇ : ∀ {c} {C₁ C₂ C₃ : Container c} {X}
- P (join′ : (C₁ ⟨∘⟩ C₂) ⊸ C₃) (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) →
- let join : ∀ {X} → ⟦ C₁ ⟧ (⟦ C₂ ⟧ X) → ⟦ C₃ ⟧ X
- join = λ {_} → ⟪ join′ ⟫⊸ ∘
- _⟨$⟩_ (Inverse.from (Composition.correct C₁ C₂)) in
- ◇ P (join xss) ↔ ◇ (◇ P) xss
-join↔◇ {C₁ = C₁} {C₂} P join xss =
- ◇ P (⟪ join ⟫⊸ xss′) ↔⟨ remove-linear P join ⟩
- ◇ P xss′ ↔⟨ sym $ flatten P xss ⟩
- ◇ (◇ P) xss ∎
- where xss′ = Inverse.from (Composition.correct C₁ C₂) ⟨$⟩ xss
+module _ {s₁ s₂ s₃ p₁ p₂ p₃}
+ {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂} {C₃ : Container s₃ p₃}
+ {x} {X : Set x} {ℓ} (P : Pred X ℓ) where
+
+ join↔◇ : (join′ : (C₁ ⟨∘⟩ C₂) ⊸ C₃) (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) →
+ let join : ∀ {X} → ⟦ C₁ ⟧ (⟦ C₂ ⟧ X) → ⟦ C₃ ⟧ X
+ join = λ {_} → ⟪ join′ ⟫⊸ ∘
+ _⟨$⟩_ (Inverse.from (Composition.correct C₁ C₂)) in
+ ◇ P (join xss) ↔ ◇ (◇ P) xss
+ join↔◇ join xss =
+ ◇ P (⟪ join ⟫⊸ xss′) ↔⟨ remove-linear P join ⟩
+ ◇ P xss′ ↔⟨ SK-sym $ flatten P xss ⟩
+ ◇ (◇ P) xss ∎
+ where xss′ = Inverse.from (Composition.correct C₁ C₂) ⟨$⟩ xss
diff --git a/src/Data/Container/Combinator.agda b/src/Data/Container/Combinator.agda
index 02f5283..5acde98 100644
--- a/src/Data/Container/Combinator.agda
+++ b/src/Data/Container/Combinator.agda
@@ -12,29 +12,31 @@ open import Data.Product as Prod hiding (Σ) renaming (_×_ to _⟨×⟩_)
open import Data.Sum renaming (_⊎_ to _⟨⊎⟩_)
open import Data.Unit.Base using (⊤)
open import Function as F hiding (id; const) renaming (_∘_ to _⟨∘⟩_)
-open import Function.Inverse using (_↔_)
+open import Function.Inverse using (_↔_; inverse)
open import Level
open import Relation.Binary.PropositionalEquality as P
- using (_≗_; refl)
+ using (_≗_; _≡_; refl)
------------------------------------------------------------------------
-- Combinators
+module _ {s p : Level} where
+
-- Identity.
-id : ∀ {c} → Container c
-id = Lift ⊤ ▷ F.const (Lift ⊤)
+ id : Container s p
+ id = Lift s ⊤ ▷ F.const (Lift p ⊤)
-- Constant.
-const : ∀ {c} → Set c → Container c
-const X = X ▷ F.const (Lift ⊥)
+ const : Set s → Container s p
+ const X = X ▷ F.const (Lift p ⊥)
-- Composition.
infixr 9 _∘_
-_∘_ : ∀ {c} → Container c → Container c → Container c
+_∘_ : ∀ {s₁ s₂ p₁ p₂} → Container s₁ p₁ → Container s₂ p₂ → Container (s₁ ⊔ s₂ ⊔ p₁) (p₁ ⊔ p₂)
C₁ ∘ C₂ = ⟦ C₁ ⟧ (Shape C₂) ▷ ◇ (Position C₂)
-- Product. (Note that, up to isomorphism, this is a special case of
@@ -42,14 +44,14 @@ C₁ ∘ C₂ = ⟦ C₁ ⟧ (Shape C₂) ▷ ◇ (Position C₂)
infixr 2 _×_
-_×_ : ∀ {c} → Container c → Container c → Container c
+_×_ : ∀ {s₁ s₂ p₁ p₂} → Container s₁ p₁ → Container s₂ p₂ → Container (s₁ ⊔ s₂) (p₁ ⊔ p₂)
C₁ × C₂ =
(Shape C₁ ⟨×⟩ Shape C₂) ▷
uncurry (λ s₁ s₂ → Position C₁ s₁ ⟨⊎⟩ Position C₂ s₂)
-- Indexed product.
-Π : ∀ {c} {I : Set c} → (I → Container c) → Container c
+Π : ∀ {i s p} {I : Set i} → (I → Container s p) → Container (i ⊔ s) (i ⊔ p)
Π C = (∀ i → Shape (C i)) ▷ λ s → ∃ λ i → Position (C i) (s i)
-- Sum. (Note that, up to isomorphism, this is a special case of
@@ -57,12 +59,12 @@ C₁ × C₂ =
infixr 1 _⊎_
-_⊎_ : ∀ {c} → Container c → Container c → Container c
+_⊎_ : ∀ {s₁ s₂ p} → Container s₁ p → Container s₂ p → Container (s₁ ⊔ s₂) p
C₁ ⊎ C₂ = (Shape C₁ ⟨⊎⟩ Shape C₂) ▷ [ Position C₁ , Position C₂ ]
-- Indexed sum.
-Σ : ∀ {c} {I : Set c} → (I → Container c) → Container c
+Σ : ∀ {i s p} {I : Set i} → (I → Container s p) → Container (i ⊔ s) p
Σ C = (∃ λ i → Shape (C i)) ▷ λ s → Position (C (proj₁ s)) (proj₂ s)
-- Constant exponentiation. (Note that this is a special case of
@@ -70,7 +72,7 @@ C₁ ⊎ C₂ = (Shape C₁ ⟨⊎⟩ Shape C₂) ▷ [ Position C₁ , Position
infix 0 const[_]⟶_
-const[_]⟶_ : ∀ {c} → Set c → Container c → Container c
+const[_]⟶_ : ∀ {i s p} → Set i → Container s p → Container (i ⊔ s) (i ⊔ p)
const[ X ]⟶ C = Π {I = X} (F.const C)
------------------------------------------------------------------------
@@ -82,26 +84,24 @@ const[ X ]⟶ C = Π {I = X} (F.const C)
module Identity where
- 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)
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct : ∀ {s p x} {X : Set x} → ⟦ id {s} {p} ⟧ X ↔ F.id X
+ correct {X = X} = inverse to from (λ _ → refl) (λ _ → refl)
+ where
+ to : ⟦ id ⟧ X → F.id X
+ to xs = proj₂ xs _
-module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
+ from : F.id X → ⟦ id ⟧ X
+ from x = (_ , λ _ → x)
- correct : ∀ {ℓ} (X : Set ℓ) {Y} → ⟦ const X ⟧ Y ↔ F.const X Y
- correct X {Y} = record
+module Constant (ext : ∀ {ℓ ℓ′} → P.Extensionality ℓ ℓ′) where
+
+ correct : ∀ {x p y} (X : Set x) {Y : Set y} → ⟦ const {x} {p ⊔ y} X ⟧ Y ↔ F.const X Y
+ correct {x} {y} X {Y} = record
{ to = P.→-to-⟶ to
; from = P.→-to-⟶ from
; inverse-of = record
{ right-inverse-of = λ _ → refl
- ; left-inverse-of =
- λ xs → P.cong (_,_ (proj₁ xs)) (ext (λ x → ⊥-elim (lower x)))
+ ; left-inverse-of = from∘to
}
}
where
@@ -111,18 +111,13 @@ module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
from : X → ⟦ const X ⟧ Y
from = < F.id , F.const (⊥-elim ∘′ lower) >
-module Composition where
+ from∘to : (x : ⟦ const X ⟧ Y) → from (to x) ≡ x
+ from∘to xs = P.cong (proj₁ xs ,_) (ext (λ x → ⊥-elim (lower x)))
- correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
- ⟦ C₁ ∘ C₂ ⟧ X ↔ (⟦ C₁ ⟧ ⟨∘⟩ ⟦ C₂ ⟧) X
- correct C₁ C₂ {X} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
+module Composition {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) where
+
+ correct : ∀ {x} {X : Set x} → ⟦ C₁ ∘ C₂ ⟧ X ↔ (⟦ C₁ ⟧ ⟨∘⟩ ⟦ C₂ ⟧) X
+ correct {X = X} = inverse to from (λ _ → refl) (λ _ → refl)
where
to : ⟦ C₁ ∘ C₂ ⟧ X → ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)
to ((s , f) , g) = (s , < f , curry g >)
@@ -130,18 +125,11 @@ module Composition where
from : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X) → ⟦ C₁ ∘ C₂ ⟧ X
from (s , f) = ((s , proj₁ ⟨∘⟩ f) , uncurry (proj₂ ⟨∘⟩ f))
-module Product (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
+module Product (ext : ∀ {ℓ ℓ′} → P.Extensionality ℓ ℓ′)
+ {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) where
- correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
- ⟦ C₁ × C₂ ⟧ X ↔ (⟦ C₁ ⟧ X ⟨×⟩ ⟦ C₂ ⟧ X)
- correct {c} C₁ C₂ {X} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct : ∀ {x} {X : Set x} → ⟦ C₁ × C₂ ⟧ X ↔ (⟦ C₁ ⟧ X ⟨×⟩ ⟦ C₂ ⟧ X)
+ correct {X = X} = inverse to from from∘to (λ _ → refl)
where
to : ⟦ C₁ × C₂ ⟧ X → ⟦ C₁ ⟧ X ⟨×⟩ ⟦ C₂ ⟧ X
to ((s₁ , s₂) , f) = ((s₁ , f ⟨∘⟩ inj₁) , (s₂ , f ⟨∘⟩ inj₂))
@@ -151,39 +139,23 @@ module Product (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
from∘to : from ⟨∘⟩ to ≗ F.id
from∘to (s , f) =
- P.cong (_,_ s) (ext {ℓ = c} [ (λ _ → refl) , (λ _ → refl) ])
+ P.cong (s ,_) (ext [ (λ _ → refl) , (λ _ → refl) ])
-module IndexedProduct where
+module IndexedProduct {i s p} {I : Set i} (Cᵢ : I → Container s p) where
- correct : ∀ {c I} (C : I → Container c) {X : Set c} →
- ⟦ Π C ⟧ X ↔ (∀ i → ⟦ C i ⟧ X)
- correct {I = I} C {X} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct : ∀ {x} {X : Set x} → ⟦ Π Cᵢ ⟧ X ↔ (∀ i → ⟦ Cᵢ i ⟧ X)
+ correct {X = X} = inverse to from (λ _ → refl) (λ _ → refl)
where
- to : ⟦ Π C ⟧ X → ∀ i → ⟦ C i ⟧ X
+ to : ⟦ Π Cᵢ ⟧ X → ∀ i → ⟦ Cᵢ i ⟧ X
to (s , f) = λ i → (s i , λ p → f (i , p))
- from : (∀ i → ⟦ C i ⟧ X) → ⟦ Π C ⟧ X
+ from : (∀ i → ⟦ Cᵢ i ⟧ X) → ⟦ Π Cᵢ ⟧ X
from f = (proj₁ ⟨∘⟩ f , uncurry (proj₂ ⟨∘⟩ f))
-module Sum where
+module Sum {s₁ s₂ p} (C₁ : Container s₁ p) (C₂ : Container s₂ p) where
- correct : ∀ {c} (C₁ C₂ : Container c) {X : Set c} →
- ⟦ C₁ ⊎ C₂ ⟧ X ↔ (⟦ C₁ ⟧ X ⟨⊎⟩ ⟦ C₂ ⟧ X)
- correct C₁ C₂ {X} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = [ (λ _ → refl) , (λ _ → refl) ]
- }
- }
+ correct : ∀ {x} {X : Set x} → ⟦ C₁ ⊎ C₂ ⟧ X ↔ (⟦ C₁ ⟧ X ⟨⊎⟩ ⟦ C₂ ⟧ X)
+ correct {X = X} = inverse to from from∘to to∘from
where
to : ⟦ C₁ ⊎ C₂ ⟧ X → ⟦ C₁ ⟧ X ⟨⊎⟩ ⟦ C₂ ⟧ X
to (inj₁ s₁ , f) = inj₁ (s₁ , f)
@@ -196,18 +168,13 @@ module Sum where
from∘to (inj₁ s₁ , f) = refl
from∘to (inj₂ s₂ , f) = refl
-module IndexedSum where
+ to∘from : to ⟨∘⟩ from ≗ F.id
+ to∘from = [ (λ _ → refl) , (λ _ → refl) ]
- correct : ∀ {c I} (C : I → Container c) {X : Set c} →
- ⟦ Σ C ⟧ X ↔ (∃ λ i → ⟦ C i ⟧ X)
- correct {I = I} C {X} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
+module IndexedSum {i s p} {I : Set i} (C : I → Container s p) where
+
+ correct : ∀ {x} {X : Set x} → ⟦ Σ C ⟧ X ↔ (∃ λ i → ⟦ C i ⟧ X)
+ correct {X = X} = inverse to from (λ _ → refl) (λ _ → refl)
where
to : ⟦ Σ C ⟧ X → ∃ λ i → ⟦ C i ⟧ X
to ((i , s) , f) = (i , (s , f))
@@ -215,8 +182,7 @@ module IndexedSum where
from : (∃ λ i → ⟦ C i ⟧ X) → ⟦ Σ C ⟧ X
from (i , (s , f)) = ((i , s) , f)
-module ConstantExponentiation where
+module ConstantExponentiation {i s p} {I : Set i} (C : Container s p) where
- correct : ∀ {c X} (C : Container c) {Y : Set c} →
- ⟦ const[ X ]⟶ C ⟧ Y ↔ (X → ⟦ C ⟧ Y)
- correct C = IndexedProduct.correct (F.const C)
+ correct : ∀ {x} {X : Set x} → ⟦ const[ I ]⟶ C ⟧ X ↔ (I → ⟦ C ⟧ X)
+ correct = IndexedProduct.correct (F.const C)
diff --git a/src/Data/Container/Core.agda b/src/Data/Container/Core.agda
new file mode 100644
index 0000000..e44a3f5
--- /dev/null
+++ b/src/Data/Container/Core.agda
@@ -0,0 +1,53 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Containers core
+------------------------------------------------------------------------
+
+module Data.Container.Core where
+
+open import Level
+open import Data.Product
+open import Function
+open import Relation.Unary using (Pred)
+
+-- Definition of Containers
+
+infix 5 _▷_
+record Container (s p : Level) : Set (suc (s ⊔ p)) where
+ constructor _▷_
+ field
+ Shape : Set s
+ Position : Shape → Set p
+open Container
+
+-- The semantics ("extension") of a container.
+
+⟦_⟧ : ∀ {s p ℓ} → Container s p → Set ℓ → Set (s ⊔ p ⊔ ℓ)
+⟦ S ▷ P ⟧ X = Σ[ s ∈ S ] (P s → X)
+
+-- Representation of container morphisms.
+
+record _⇒_ {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂)
+ : Set (s₁ ⊔ s₂ ⊔ p₁ ⊔ p₂) where
+ constructor _▷_
+ field
+ shape : Shape C₁ → Shape C₂
+ position : ∀ {s} → Position C₂ (shape s) → Position C₁ s
+open _⇒_ public
+
+-- Interpretation of _⇒_.
+
+⟪_⟫ : ∀ {s₁ s₂ p₁ p₂ x} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂} →
+ C₁ ⇒ C₂ → {X : Set x} → ⟦ C₁ ⟧ X → ⟦ C₂ ⟧ X
+⟪ m ⟫ = map (shape m) (_∘ position m)
+
+-- All and Any
+
+module _ {s p} {C : Container s p} {x} {X : Set x} where
+
+ □ : ∀ {ℓ} → Pred X ℓ → Pred (⟦ C ⟧ X) (p ⊔ ℓ)
+ □ P (s , f) = ∀ p → P (f p)
+
+ ◇ : ∀ {ℓ} → Pred X ℓ → Pred (⟦ C ⟧ X) (p ⊔ ℓ)
+ ◇ P (s , f) = ∃ λ p → P (f p)
diff --git a/src/Data/Container/FreeMonad.agda b/src/Data/Container/FreeMonad.agda
index f3e8a25..3c19df3 100644
--- a/src/Data/Container/FreeMonad.agda
+++ b/src/Data/Container/FreeMonad.agda
@@ -7,9 +7,7 @@
module Data.Container.FreeMonad where
open import Level
-open import Function using (_∘_)
-open import Data.Empty using (⊥-elim)
-open import Data.Sum using (inj₁; inj₂)
+open import Data.Sum using (inj₁; inj₂ ; [_,_]′)
open import Data.Product
open import Data.Container
open import Data.Container.Combinator using (const; _⊎_)
@@ -35,21 +33,23 @@ infix 1 _⋆_
-- up in a leaf (element of the set) -- hence the Kleene star notation
-- (the type can be read as a regular expression).
-_⋆C_ : ∀ {c} → Container c → Set c → Container c
+_⋆C_ : ∀ {x s p} → Container s p → Set x → Container (s ⊔ x) p
C ⋆C X = const X ⊎ C
-_⋆_ : ∀ {c} → Container c → Set c → Set c
+_⋆_ : ∀ {x s p} → Container s p → Set x → Set (x ⊔ s ⊔ p)
C ⋆ X = μ (C ⋆C X)
-do : ∀ {c} {C : Container c} {X} → ⟦ C ⟧ (C ⋆ X) → C ⋆ X
-do (s , k) = sup (inj₂ s) k
+module _ {s p} {C : Container s p} where
-rawMonad : ∀ {c} {C : Container c} → RawMonad (_⋆_ C)
-rawMonad = record { return = return; _>>=_ = _>>=_ }
- where
- return : ∀ {c} {C : Container c} {X} → X → C ⋆ X
- return x = sup (inj₁ x) (⊥-elim ∘ lower)
+ inn : ∀ {x} {X : Set x} → ⟦ C ⟧ (C ⋆ X) → C ⋆ X
+ inn (s , f) = sup (inj₂ s , f)
- _>>=_ : ∀ {c} {C : Container c} {X Y} → C ⋆ X → (X → C ⋆ Y) → C ⋆ Y
- sup (inj₁ x) _ >>= k = k x
- sup (inj₂ s) f >>= k = do (s , λ p → f p >>= k)
+ rawMonad : ∀ {x} → RawMonad {s ⊔ p ⊔ x} (C ⋆_)
+ rawMonad = record { return = return; _>>=_ = _>>=_ }
+ where
+ return : ∀ {X} → X → C ⋆ X
+ return x = sup (inj₁ x , λ ())
+
+ _>>=_ : ∀ {X Y} → C ⋆ X → (X → C ⋆ Y) → C ⋆ Y
+ sup (inj₁ x , _) >>= k = k x
+ sup (inj₂ s , f) >>= k = inn (s , λ p → f p >>= k)
diff --git a/src/Data/Container/Indexed.agda b/src/Data/Container/Indexed.agda
index 184e46e..78ca953 100644
--- a/src/Data/Container/Indexed.agda
+++ b/src/Data/Container/Indexed.agda
@@ -11,17 +11,17 @@
module Data.Container.Indexed where
open import Level
+open import Codata.Musical.M.Indexed
+open import Data.Product as Prod hiding (map)
+open import Data.W.Indexed
open import Function renaming (id to ⟨id⟩; _∘_ to _⟨∘⟩_)
open import Function.Equality using (_⟨$⟩_)
open import Function.Inverse using (_↔_; module Inverse)
-open import Data.Product as Prod hiding (map)
open import Relation.Unary using (Pred; _⊆_)
open import Relation.Binary as B using (Preorder; module Preorder)
open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_; refl)
open import Relation.Binary.HeterogeneousEquality as H using (_≅_; refl)
-open import Relation.Binary.Indexed
-open import Data.W.Indexed
-open import Data.M.Indexed
+open import Relation.Binary.Indexed.Heterogeneous hiding (Rel; REL)
------------------------------------------------------------------------
@@ -45,7 +45,7 @@ I ▷ O = Container I O zero zero
-- Equality, parametrised on an underlying relation.
Eq : ∀ {i o c r ℓ} {I : Set i} {O : Set o} (C : Container I O c r)
- (X Y : Pred I ℓ) → REL X Y ℓ → REL (⟦ C ⟧ X) (⟦ C ⟧ Y) _
+ (X Y : Pred I ℓ) → IREL X Y ℓ → IREL (⟦ C ⟧ X) (⟦ C ⟧ Y) _
Eq C _ _ _≈_ {o₁} {o₂} (c , k) (c′ , k′) =
o₁ ≡ o₂ × c ≅ c′ × (∀ r r′ → r ≅ r′ → k r ≈ k′ r′)
@@ -62,7 +62,7 @@ private
H.cong (_,_ c) (ext (λ _ → refl) (λ r → k≈k′ r r refl))
setoid : ∀ {i o c r s} {I : Set i} {O : Set o} →
- Container I O c r → Setoid I s _ → Setoid O _ _
+ Container I O c r → IndexedSetoid I s _ → IndexedSetoid O _ _
setoid C X = record
{ Carrier = ⟦ C ⟧ X.Carrier
; _≈_ = _≈_
@@ -73,9 +73,9 @@ setoid C X = record
}
}
where
- module X = Setoid X
+ module X = IndexedSetoid X
- _≈_ : Rel (⟦ C ⟧ X.Carrier) _
+ _≈_ : IRel (⟦ C ⟧ X.Carrier) _
_≈_ = Eq C X.Carrier X.Carrier X._≈_
sym : Symmetric (⟦ C ⟧ X.Carrier) _≈_
@@ -100,19 +100,19 @@ map _ f = Prod.map ⟨id⟩ (λ g → f ⟨∘⟩ g)
module Map where
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
+ (X : IndexedSetoid I s _) → let module X = IndexedSetoid 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)
+ identity C X = IndexedSetoid.refl (setoid C X)
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 s _) → let module Z = Setoid Z in
+ (Z : IndexedSetoid I s _) → let module Z = IndexedSetoid Z in
{f : Y ⊆ Z.Carrier} {g : X ⊆ Y} {o : O} {xs : ⟦ C ⟧ X o} →
Eq C Z.Carrier Z.Carrier Z._≈_
(map C {Y} f (map C {X} g xs))
(map C {X} (f ⟨∘⟩ g) xs)
- composition C Z = Setoid.refl (setoid C Z)
+ composition C Z = IndexedSetoid.refl (setoid C Z)
------------------------------------------------------------------------
-- Container morphisms
@@ -190,7 +190,7 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where
Natural : ∀ {ℓ} {C₁ C₂ : Container I O c r} →
((X : Pred I ℓ) → ⟦ C₁ ⟧ X ⊆ ⟦ C₂ ⟧ X) → Set _
Natural {C₁ = C₁} {C₂} m =
- ∀ {X} Y → let module Y = Setoid Y in (f : X ⊆ Y.Carrier) →
+ ∀ {X} Y → let module Y = IndexedSetoid Y in (f : X ⊆ Y.Carrier) →
∀ {o} (xs : ⟦ C₁ ⟧ X o) →
Eq C₂ Y.Carrier Y.Carrier Y._≈_
(m Y.Carrier $ map C₁ {X} f xs) (map C₂ {X} f $ m X xs)
@@ -206,7 +206,7 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where
natural : ∀ {ℓ} (C₁ C₂ : Container I O c r) (m : C₁ ⇒ C₂) → Natural {ℓ} ⟪ m ⟫
natural _ _ m {X} Y f _ = refl , refl , λ { r .r refl → lemma (coherent m) }
where
- module Y = Setoid Y
+ module Y = IndexedSetoid Y
lemma : ∀ {i j} (eq : i ≡ j) {x} →
P.subst Y.Carrier eq (f x) Y.≈ f (P.subst X eq x)
@@ -216,13 +216,13 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where
-- morphisms.
complete : ∀ {C₁ C₂ : Container I O c r} (nt : NT C₁ C₂) →
- ∃ λ m → (X : Setoid I _ _) →
- let module X = Setoid X in
+ ∃ λ m → (X : IndexedSetoid I _ _) →
+ let module X = IndexedSetoid X in
∀ {o} (xs : ⟦ C₁ ⟧ X.Carrier o) →
Eq C₂ X.Carrier X.Carrier X._≈_
(proj₁ nt X.Carrier xs) (⟪ m ⟫ X.Carrier {o} xs)
complete {C₁} {C₂} (nt , nat) = m , (λ X xs → nat X
- (λ { (r , eq) → P.subst (Setoid.Carrier X) eq (proj₂ xs r) })
+ (λ { (r , eq) → P.subst (IndexedSetoid.Carrier X) eq (proj₂ xs r) })
(proj₁ xs , (λ r → r , refl)))
where
@@ -267,8 +267,8 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where
id-correct _ = refl
∘-correct : {C₁ C₂ C₃ : Container I O c r}
- (f : C₂ ⇒ C₃) (g : C₁ ⇒ C₂) (X : Setoid I (c ⊔ r) _) →
- let module X = Setoid X in
+ (f : C₂ ⇒ C₃) (g : C₁ ⇒ C₂) (X : IndexedSetoid I (c ⊔ r) _) →
+ let module X = IndexedSetoid X in
∀ {o} {xs : ⟦ C₁ ⟧ X.Carrier o} →
Eq C₃ X.Carrier X.Carrier X._≈_
(⟪ f ∘ g ⟫ X.Carrier xs)
@@ -276,7 +276,7 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where
∘-correct f g X = refl , refl , λ { r .r refl → lemma (coherent g)
(coherent f) }
where
- module X = Setoid X
+ module X = IndexedSetoid X
lemma : ∀ {i j k} (eq₁ : i ≡ j) (eq₂ : j ≡ k) {x} →
P.subst X.Carrier (P.trans eq₁ eq₂) x
@@ -347,5 +347,5 @@ module _ {i o c r ℓ₁ ℓ₂} {I : Set i} {O : Set o} (C : Container I O c r)
infix 4 _∈_
_∈_ : ∀ {i o c r ℓ} {I : Set i} {O : Set o}
- {C : Container I O c r} {X : Pred I (i ⊔ ℓ)} → REL X (⟦ C ⟧ X) _
-_∈_ {C = C} {X} x xs = ◇ C {X = X} (_≅_ x) (, xs)
+ {C : Container I O c r} {X : Pred I (i ⊔ ℓ)} → IREL X (⟦ C ⟧ X) _
+_∈_ {C = C} {X} x xs = ◇ C {X = X} (_≅_ x) (-, xs)
diff --git a/src/Data/Container/Indexed/Combinator.agda b/src/Data/Container/Indexed/Combinator.agda
index bffc774..95d5e38 100644
--- a/src/Data/Container/Indexed/Combinator.agda
+++ b/src/Data/Container/Indexed/Combinator.agda
@@ -13,7 +13,7 @@ open import Data.Unit.Base using (⊤)
open import Data.Product as Prod hiding (Σ) renaming (_×_ to _⟨×⟩_)
open import Data.Sum renaming (_⊎_ to _⟨⊎⟩_)
open import Function as F hiding (id; const) renaming (_∘_ to _⟨∘⟩_)
-open import Function.Inverse using (_↔̇_)
+open import Function.Inverse using (_↔̇_; inverse)
open import Relation.Unary using (Pred; _⊆_; _∪_; _∩_; ⋃; ⋂)
renaming (_⟨×⟩_ to _⟪×⟫_; _⟨⊙⟩_ to _⟪⊙⟫_; _⟨⊎⟩_ to _⟪⊎⟫_)
open import Relation.Binary.PropositionalEquality as P
@@ -26,13 +26,13 @@ open import Relation.Binary.PropositionalEquality as P
-- Identity.
id : ∀ {o c r} {O : Set o} → Container O O c r
-id = F.const (Lift ⊤) ◃ (λ _ → Lift ⊤) / (λ {o} _ _ → o)
+id = F.const (Lift _ ⊤) ◃ (λ _ → Lift _ ⊤) / (λ {o} _ _ → o)
-- Constant.
const : ∀ {i o c r} {I : Set i} {O : Set o} →
Pred O c → Container I O c r
-const X = X ◃ (λ _ → Lift ⊥) / λ _ → ⊥-elim ⟨∘⟩ lower
+const X = X ◃ (λ _ → Lift _ ⊥) / λ _ → ⊥-elim ⟨∘⟩ lower
-- Duality.
@@ -137,14 +137,13 @@ module Identity where
correct : ∀ {o ℓ c r} {O : Set o} {X : Pred O ℓ} →
⟦ id {c = c}{r} ⟧ X ↔̇ F.id X
- correct = record
- { to = P.→-to-⟶ λ xs → proj₂ xs _
- ; from = P.→-to-⟶ λ x → (_ , λ _ → x)
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct {X = X} = inverse to from (λ _ → refl) (λ _ → refl)
+ where
+ to : ∀ {x} → ⟦ id ⟧ X x → F.id X x
+ to xs = proj₂ xs _
+
+ from : ∀ {x} → F.id X x → ⟦ id ⟧ X x
+ from x = (_ , λ _ → x)
module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
@@ -155,8 +154,7 @@ module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
; from = P.→-to-⟶ from
; inverse-of = record
{ right-inverse-of = λ _ → refl
- ; left-inverse-of =
- λ xs → P.cong (_,_ (proj₁ xs)) (ext (⊥-elim ⟨∘⟩ lower))
+ ; left-inverse-of = to∘from
}
}
where
@@ -166,33 +164,23 @@ module Constant (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
from : X ⊆ ⟦ const X ⟧ Y
from = < F.id , F.const (⊥-elim ⟨∘⟩ lower) >
+ to∘from : _
+ to∘from xs = P.cong (proj₁ xs ,_) (ext (⊥-elim ⟨∘⟩ lower))
+
module Duality where
correct : ∀ {i o c r ℓ} {I : Set i} {O : Set o}
(C : Container I O c r) (X : Pred I ℓ) →
⟦ C ^⊥ ⟧ X ↔̇ (λ o → (c : Command C o) → ∃ λ r → X (next C c r))
- correct C X = record
- { to = P.→-to-⟶ λ { (f , g) → < f , g > }
- ; from = P.→-to-⟶ λ f → proj₁ ⟨∘⟩ f , proj₂ ⟨∘⟩ f
- ; inverse-of = record
- { left-inverse-of = λ { (_ , _) → refl }
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct C X = inverse (λ { (f , g) → < f , g > }) (λ f → proj₁ ⟨∘⟩ f , proj₂ ⟨∘⟩ f)
+ (λ _ → refl) (λ _ → refl)
module Composition where
correct : ∀ {i j k ℓ c r} {I : Set i} {J : Set j} {K : Set k}
(C₁ : Container J K c r) (C₂ : Container I J c r) →
{X : Pred I ℓ} → ⟦ C₁ ∘ C₂ ⟧ X ↔̇ (⟦ C₁ ⟧ ⟨∘⟩ ⟦ C₂ ⟧) X
- correct C₁ C₂ {X} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct C₁ C₂ {X} = inverse to from (λ _ → refl) (λ _ → refl)
where
to : ⟦ C₁ ∘ C₂ ⟧ X ⊆ ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)
to ((c , f) , g) = (c , < f , curry g >)
@@ -205,14 +193,7 @@ module Product (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
correct : ∀ {i o c r} {I : Set i} {O : Set o}
(C₁ C₂ : Container I O c r) {X} →
⟦ C₁ × C₂ ⟧ X ↔̇ (⟦ C₁ ⟧ X ∩ ⟦ C₂ ⟧ X)
- correct C₁ C₂ {X} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct C₁ C₂ {X} = inverse to from from∘to (λ _ → refl)
where
to : ⟦ C₁ × C₂ ⟧ X ⊆ ⟦ C₁ ⟧ X ∩ ⟦ C₂ ⟧ X
to ((c₁ , c₂) , k) = ((c₁ , k ⟨∘⟩ inj₁) , (c₂ , k ⟨∘⟩ inj₂))
@@ -222,21 +203,14 @@ module Product (ext : ∀ {ℓ} → P.Extensionality ℓ ℓ) where
from∘to : from ⟨∘⟩ to ≗ F.id
from∘to (c , _) =
- P.cong (_,_ c) (ext [ (λ _ → refl) , (λ _ → refl) ])
+ P.cong (c ,_) (ext [ (λ _ → refl) , (λ _ → refl) ])
module IndexedProduct where
correct : ∀ {x i o c r ℓ} {X : Set x} {I : Set i} {O : Set o}
(C : X → Container I O c r) {Y : Pred I ℓ} →
⟦ Π C ⟧ Y ↔̇ ⋂[ x ∶ X ] ⟦ C x ⟧ Y
- correct {X = X} C {Y} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct {X = X} C {Y} = inverse to from (λ _ → refl) (λ _ → refl)
where
to : ⟦ Π C ⟧ Y ⊆ ⋂[ x ∶ X ] ⟦ C x ⟧ Y
to (c , k) = λ x → (c x , λ r → k (x , r))
@@ -249,14 +223,7 @@ module Sum where
correct : ∀ {i o c r ℓ} {I : Set i} {O : Set o}
(C₁ C₂ : Container I O c r) {X : Pred I ℓ} →
⟦ C₁ ⊎ C₂ ⟧ X ↔̇ (⟦ C₁ ⟧ X ∪ ⟦ C₂ ⟧ X)
- correct C₁ C₂ {X} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = from∘to
- ; right-inverse-of = [ (λ _ → refl) , (λ _ → refl) ]
- }
- }
+ correct C₁ C₂ {X} = inverse to from from∘to to∘from
where
to : ⟦ C₁ ⊎ C₂ ⟧ X ⊆ ⟦ C₁ ⟧ X ∪ ⟦ C₂ ⟧ X
to (inj₁ c₁ , k) = inj₁ (c₁ , k)
@@ -269,19 +236,15 @@ module Sum where
from∘to (inj₁ _ , _) = refl
from∘to (inj₂ _ , _) = refl
+ to∘from : to ⟨∘⟩ from ≗ F.id
+ to∘from = [ (λ _ → refl) , (λ _ → refl) ]
+
module IndexedSum where
correct : ∀ {x i o c r ℓ} {X : Set x} {I : Set i} {O : Set o}
(C : X → Container I O c r) {Y : Pred I ℓ} →
⟦ Σ C ⟧ Y ↔̇ ⋃[ x ∶ X ] ⟦ C x ⟧ Y
- correct {X = X} C {Y} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ from
- ; inverse-of = record
- { left-inverse-of = λ _ → refl
- ; right-inverse-of = λ _ → refl
- }
- }
+ correct {X = X} C {Y} = inverse to from (λ _ → refl) (λ _ → refl)
where
to : ⟦ Σ C ⟧ Y ⊆ ⋃[ x ∶ X ] ⟦ C x ⟧ Y
to ((x , c) , k) = (x , (c , k))
diff --git a/src/Data/Container/Indexed/FreeMonad.agda b/src/Data/Container/Indexed/FreeMonad.agda
index b94510c..3a1a794 100644
--- a/src/Data/Container/Indexed/FreeMonad.agda
+++ b/src/Data/Container/Indexed/FreeMonad.agda
@@ -34,9 +34,9 @@ C ⋆ X = μ (C ⋆C X)
pattern returnP x = (inj₁ x , _)
pattern doP c k = (inj₂ c , k)
-do : ∀ {ℓ} {O : Set ℓ} {C : Container O O ℓ ℓ} {X} →
- ⟦ C ⟧ (C ⋆ X) ⊆ C ⋆ X
-do (c , k) = sup (doP c k)
+inn : ∀ {ℓ} {O : Set ℓ} {C : Container O O ℓ ℓ} {X} →
+ ⟦ C ⟧ (C ⋆ X) ⊆ C ⋆ X
+inn (c , k) = sup (doP c k)
rawPMonad : ∀ {ℓ} {O : Set ℓ} {C : Container O O ℓ ℓ} →
RawPMonad {ℓ = ℓ} (_⋆_ C)
@@ -50,17 +50,17 @@ rawPMonad {C = C} = record
_=<<_ : ∀ {X Y} → X ⊆ C ⋆ Y → C ⋆ X ⊆ C ⋆ Y
f =<< sup (returnP x) = f x
- f =<< sup (doP c k) = do (c , λ r → f =<< k r)
+ f =<< sup (doP c k) = inn (c , λ r → f =<< k r)
leaf : ∀ {ℓ} {O : Set ℓ} {C : Container O O ℓ ℓ} {X : Pred O ℓ} →
⟦ C ⟧ X ⊆ C ⋆ X
-leaf (c , k) = do (c , return? ∘ k)
+leaf (c , k) = inn (c , return? ∘ k)
where
open RawPMonad rawPMonad
generic : ∀ {ℓ} {O : Set ℓ} {C : Container O O ℓ ℓ} {o}
(c : Command C o) →
o ∈ C ⋆ (⋃[ r ∶ Response C c ] { next C c r })
-generic c = do (c , λ r → return? (r , refl))
+generic c = inn (c , λ r → return? (r , refl))
where
open RawPMonad rawPMonad
diff --git a/src/Data/Covec.agda b/src/Data/Covec.agda
deleted file mode 100644
index 529b0e8..0000000
--- a/src/Data/Covec.agda
+++ /dev/null
@@ -1,155 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- Coinductive vectors
-------------------------------------------------------------------------
-
-module Data.Covec where
-
-open import Coinduction
-open import Data.Nat.Base using (ℕ; zero; suc)
-open import Data.Conat as Coℕ using (Coℕ; zero; suc; _+_)
-open import Data.Cofin using (Cofin; zero; suc)
-open import Data.Vec using (Vec; []; _∷_)
-open import Data.Colist as Colist using (Colist; []; _∷_)
-open import Data.Product using (_,_)
-open import Relation.Binary
-
-------------------------------------------------------------------------
--- The type
-
-infixr 5 _∷_
-
-data Covec (A : Set) : Coℕ → Set where
- [] : Covec A zero
- _∷_ : ∀ {n} (x : A) (xs : ∞ (Covec A (♭ n))) → Covec A (suc n)
-
-------------------------------------------------------------------------
--- Some operations
-
-map : ∀ {A B n} → (A → B) → Covec A n → Covec B n
-map f [] = []
-map f (x ∷ xs) = f x ∷ ♯ map f (♭ xs)
-
-fromVec : ∀ {A n} → Vec A n → Covec A (Coℕ.fromℕ n)
-fromVec [] = []
-fromVec (x ∷ xs) = x ∷ ♯ fromVec xs
-
-fromColist : ∀ {A} (xs : Colist A) → Covec A (Colist.length xs)
-fromColist [] = []
-fromColist (x ∷ xs) = x ∷ ♯ fromColist (♭ xs)
-
-take : ∀ {A} m {n} → Covec A (m + n) → Covec A m
-take zero xs = []
-take (suc n) (x ∷ xs) = x ∷ ♯ take (♭ n) (♭ xs)
-
-drop : ∀ {A} m {n} → Covec A (Coℕ.fromℕ m + n) → Covec A n
-drop zero xs = xs
-drop (suc n) (x ∷ xs) = drop n (♭ xs)
-
-replicate : ∀ {A} n → A → Covec A n
-replicate zero x = []
-replicate (suc n) x = x ∷ ♯ replicate (♭ n) x
-
-lookup : ∀ {A n} → Cofin n → Covec A n → A
-lookup zero (x ∷ xs) = x
-lookup (suc n) (x ∷ xs) = lookup n (♭ xs)
-
-infixr 5 _++_
-
-_++_ : ∀ {A m n} → Covec A m → Covec A n → Covec A (m + n)
-[] ++ ys = ys
-(x ∷ xs) ++ ys = x ∷ ♯ (♭ xs ++ ys)
-
-[_] : ∀ {A} → A → Covec A (suc (♯ zero))
-[ x ] = x ∷ ♯ []
-
-------------------------------------------------------------------------
--- Equality and other relations
-
--- xs ≈ ys means that xs and ys are equal.
-
-infix 4 _≈_
-
-data _≈_ {A} : ∀ {n} (xs ys : Covec A n) → Set where
- [] : [] ≈ []
- _∷_ : ∀ {n} x {xs ys}
- (xs≈ : ∞ (♭ xs ≈ ♭ ys)) → _≈_ {n = suc n} (x ∷ xs) (x ∷ ys)
-
--- x ∈ xs means that x is a member of xs.
-
-infix 4 _∈_
-
-data _∈_ {A} : ∀ {n} → A → Covec A n → Set where
- here : ∀ {n x } {xs} → _∈_ {n = suc n} x (x ∷ xs)
- there : ∀ {n x y} {xs} (x∈xs : x ∈ ♭ xs) → _∈_ {n = suc n} x (y ∷ xs)
-
--- xs ⊑ ys means that xs is a prefix of ys.
-
-infix 4 _⊑_
-
-data _⊑_ {A} : ∀ {m n} → Covec A m → Covec A n → Set where
- [] : ∀ {n} {ys : Covec A n} → [] ⊑ ys
- _∷_ : ∀ {m n} x {xs ys} (p : ∞ (♭ xs ⊑ ♭ ys)) →
- _⊑_ {m = suc m} {suc n} (x ∷ xs) (x ∷ ys)
-
-------------------------------------------------------------------------
--- Some proofs
-
-setoid : Set → Coℕ → Setoid _ _
-setoid A n = record
- { Carrier = Covec A n
- ; _≈_ = _≈_
- ; isEquivalence = record
- { refl = refl
- ; sym = sym
- ; trans = trans
- }
- }
- where
- refl : ∀ {A n} → Reflexive (_≈_ {A} {n})
- refl {x = []} = []
- refl {x = x ∷ xs} = x ∷ ♯ refl
-
- sym : ∀ {A n} → Symmetric (_≈_ {A} {n})
- sym [] = []
- sym (x ∷ xs≈) = x ∷ ♯ sym (♭ xs≈)
-
- trans : ∀ {A n} → Transitive (_≈_ {A} {n})
- trans [] [] = []
- trans (x ∷ xs≈) (.x ∷ ys≈) = x ∷ ♯ trans (♭ xs≈) (♭ ys≈)
-
-poset : Set → Coℕ → Poset _ _ _
-poset A n = record
- { Carrier = Covec A n
- ; _≈_ = _≈_
- ; _≤_ = _⊑_
- ; isPartialOrder = record
- { isPreorder = record
- { isEquivalence = Setoid.isEquivalence (setoid A n)
- ; reflexive = reflexive
- ; trans = trans
- }
- ; antisym = antisym
- }
- }
- where
- reflexive : ∀ {A n} → _≈_ {A} {n} ⇒ _⊑_
- reflexive [] = []
- reflexive (x ∷ xs≈) = x ∷ ♯ reflexive (♭ xs≈)
-
- trans : ∀ {A n} → Transitive (_⊑_ {A} {n})
- trans [] _ = []
- trans (x ∷ xs≈) (.x ∷ ys≈) = x ∷ ♯ trans (♭ xs≈) (♭ ys≈)
-
- antisym : ∀ {A n} → Antisymmetric (_≈_ {A} {n}) _⊑_
- antisym [] [] = []
- antisym (x ∷ p₁) (.x ∷ p₂) = x ∷ ♯ antisym (♭ p₁) (♭ p₂)
-
-map-cong : ∀ {A B n} (f : A → B) → _≈_ {n = n} =[ map f ]⇒ _≈_
-map-cong f [] = []
-map-cong f (x ∷ xs≈) = f x ∷ ♯ map-cong f (♭ xs≈)
-
-take-⊑ : ∀ {A} m {n} (xs : Covec A (m + n)) → take m xs ⊑ xs
-take-⊑ zero xs = []
-take-⊑ (suc n) (x ∷ xs) = x ∷ ♯ take-⊑ (♭ n) (♭ xs)
diff --git a/src/Data/Digit.agda b/src/Data/Digit.agda
index 8ad6604..4ba062b 100644
--- a/src/Data/Digit.agda
+++ b/src/Data/Digit.agda
@@ -8,7 +8,7 @@ module Data.Digit where
open import Data.Nat using (ℕ; zero; suc; pred; _+_; _*_; _≤?_; _≤′_)
open import Data.Nat.Properties
-open SemiringSolver
+open import Data.Nat.Solver
open import Data.Fin as Fin using (Fin; zero; suc; toℕ)
open import Data.Char using (Char)
open import Data.List.Base
@@ -22,6 +22,8 @@ open import Relation.Binary using (Decidable)
open import Relation.Binary.PropositionalEquality as P using (_≡_; refl)
open import Function
+open +-*-Solver
+
------------------------------------------------------------------------
-- Digits
@@ -110,4 +112,3 @@ toDigits (suc (suc k)) n = <′-rec Pred helper n
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/Fin.agda b/src/Data/Fin.agda
index 12e3950..35ff551 100644
--- a/src/Data/Fin.agda
+++ b/src/Data/Fin.agda
@@ -4,240 +4,15 @@
-- Finite sets
------------------------------------------------------------------------
--- Note that elements of Fin n can be seen as natural numbers in the
--- set {m | m < n}. The notation "m" in comments below refers to this
--- natural number view.
-
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∸_
- ; _≤_ to _N≤_; _≥_ to _N≥_; _<_ to _N<_; _≤?_ to _N≤?_)
-open import Function
-import Level
-open import Relation.Nullary
-open import Relation.Nullary.Decidable
-open import Relation.Binary
-open import Relation.Binary.PropositionalEquality
- using (_≡_; _≢_; refl; cong)
-
-------------------------------------------------------------------------
--- Types
-
--- Fin n is a type with n elements.
-
-data Fin : ℕ → Set where
- zero : {n : ℕ} → Fin (suc n)
- suc : {n : ℕ} (i : Fin n) → Fin (suc n)
-
--- A conversion: toℕ "n" = n.
-
-toℕ : ∀ {n} → Fin n → ℕ
-toℕ zero = 0
-toℕ (suc i) = suc (toℕ i)
-
--- A Fin-indexed variant of Fin.
-
-Fin′ : ∀ {n} → Fin n → Set
-Fin′ i = Fin (toℕ i)
-
------------------------------------------------------------------------
--- Conversions
-
--- toℕ is defined above.
-
--- fromℕ n = "n".
-
-fromℕ : (n : ℕ) → Fin (suc n)
-fromℕ zero = zero
-fromℕ (suc n) = suc (fromℕ n)
-
--- fromℕ≤ {m} _ = "m".
-
-fromℕ≤ : ∀ {m n} → m N< n → Fin n
-fromℕ≤ (Nat.s≤s Nat.z≤n) = zero
-fromℕ≤ (Nat.s≤s (Nat.s≤s m≤n)) = suc (fromℕ≤ (Nat.s≤s m≤n))
-
--- fromℕ≤″ m _ = "m".
+-- Publicly re-export the contents of the base module
-fromℕ≤″ : ∀ m {n} → m Nat.<″ n → Fin n
-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".
-
-infix 10 #_
-
-#_ : ∀ m {n} {m<n : True (suc m N≤? n)} → Fin n
-#_ _ {m<n = m<n} = fromℕ≤ (toWitness m<n)
-
--- raise m "n" = "m + n".
-
-raise : ∀ {m} n → Fin m → Fin (n N+ m)
-raise zero i = i
-raise (suc n) i = suc (raise n i)
-
--- reduce≥ "m + n" _ = "n".
-
-reduce≥ : ∀ {m n} (i : Fin (m N+ n)) (i≥m : toℕ i N≥ m) → Fin n
-reduce≥ {zero} i i≥m = i
-reduce≥ {suc m} zero ()
-reduce≥ {suc m} (suc i) (s≤s i≥m) = reduce≥ i i≥m
-
--- inject⋆ m "n" = "n".
-
-inject : ∀ {n} {i : Fin n} → Fin′ i → Fin n
-inject {i = zero} ()
-inject {i = suc i} zero = zero
-inject {i = suc i} (suc j) = suc (inject j)
-
-inject! : ∀ {n} {i : Fin (suc n)} → Fin′ i → Fin n
-inject! {n = zero} {i = suc ()} _
-inject! {i = zero} ()
-inject! {n = suc _} {i = suc _} zero = zero
-inject! {n = suc _} {i = suc _} (suc j) = suc (inject! j)
-
-inject+ : ∀ {m} n → Fin m → Fin (m N+ n)
-inject+ n zero = zero
-inject+ n (suc i) = suc (inject+ n i)
-
-inject₁ : ∀ {m} → Fin m → Fin (suc m)
-inject₁ zero = zero
-inject₁ (suc i) = suc (inject₁ i)
-
-inject≤ : ∀ {m n} → Fin m → m N≤ n → Fin n
-inject≤ zero (Nat.s≤s le) = zero
-inject≤ (suc i) (Nat.s≤s le) = suc (inject≤ i le)
-
--- A strengthening injection into the minimal Fin fibre.
-strengthen : ∀ {n} (i : Fin n) → Fin′ (suc i)
-strengthen zero = zero
-strengthen (suc i) = suc (strengthen i)
+open import Data.Fin.Base public
------------------------------------------------------------------------
--- Operations
-
--- Folds.
-
-fold : ∀ (T : ℕ → Set) {m} →
- (∀ {n} → T n → T (suc n)) →
- (∀ {n} → T (suc n)) →
- Fin m → T m
-fold T f x zero = x
-fold T f x (suc i) = f (fold T f x i)
-
-fold′ : ∀ {n t} (T : Fin (suc n) → Set t) →
- (∀ i → T (inject₁ i) → T (suc i)) →
- T zero →
- ∀ i → T i
-fold′ T f x zero = x
-fold′ {n = zero} T f x (suc ())
-fold′ {n = suc n} T f x (suc i) =
- f i (fold′ (T ∘ inject₁) (f ∘ inject₁) x i)
-
--- Lifts functions.
-
-lift : ∀ {m n} k → (Fin m → Fin n) → Fin (k N+ m) → Fin (k N+ n)
-lift zero f i = f i
-lift (suc k) f zero = zero
-lift (suc k) f (suc i) = suc (lift k f i)
-
--- "m" + "n" = "m + n".
-
-infixl 6 _+_
-
-_+_ : ∀ {m n} (i : Fin m) (j : Fin n) → Fin (toℕ i N+ n)
-zero + j = j
-suc i + j = suc (i + j)
-
--- "m" - "n" = "m ∸ n".
-
-infixl 6 _-_
-
-_-_ : ∀ {m} (i : Fin m) (j : Fin′ (suc i)) → Fin (m N∸ toℕ j)
-i - zero = i
-zero - suc ()
-suc i - suc j = i - j
-
--- m ℕ- "n" = "m ∸ n".
-
-infixl 6 _ℕ-_
-
-_ℕ-_ : (n : ℕ) (j : Fin (suc n)) → Fin (suc n N∸ toℕ j)
-n ℕ- zero = fromℕ n
-zero ℕ- suc ()
-suc n ℕ- suc i = n ℕ- i
-
--- m ℕ-ℕ "n" = m ∸ n.
-
-infixl 6 _ℕ-ℕ_
-
-_ℕ-ℕ_ : (n : ℕ) → Fin (suc n) → ℕ
-n ℕ-ℕ zero = n
-zero ℕ-ℕ suc ()
-suc n ℕ-ℕ suc i = n ℕ-ℕ i
-
--- pred "n" = "pred n".
-
-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
-
-infix 4 _≤_ _<_
-
-_≤_ : ∀ {n} → Rel (Fin n) Level.zero
-_≤_ = _N≤_ on toℕ
-
-_≤?_ : ∀ {n} → (a : Fin n) → (b : Fin n) → Dec (a ≤ b)
-a ≤? b = toℕ a N≤? toℕ b
-
-_<_ : ∀ {n} → Rel (Fin n) Level.zero
-_<_ = _N<_ on toℕ
-
-data _≺_ : ℕ → ℕ → Set where
- _≻toℕ_ : ∀ n (i : Fin n) → toℕ i ≺ n
-
--- An ordering view.
-
-data Ordering {n : ℕ} : Fin n → Fin n → Set where
- less : ∀ greatest (least : Fin′ greatest) →
- Ordering (inject least) greatest
- equal : ∀ i → Ordering i i
- greater : ∀ greatest (least : Fin′ greatest) →
- Ordering greatest (inject least)
+-- Publicly re-export queries
-compare : ∀ {n} (i j : Fin n) → Ordering i j
-compare zero zero = equal zero
-compare zero (suc j) = less (suc j) zero
-compare (suc i) zero = greater (suc i) zero
-compare (suc i) (suc j) with compare i j
-compare (suc .(inject least)) (suc .greatest) | less greatest least =
- less (suc greatest) (suc least)
-compare (suc .greatest) (suc .(inject least)) | greater greatest least =
- greater (suc greatest) (suc least)
-compare (suc .i) (suc .i) | equal i = equal (suc i)
+open import Data.Fin.Properties public
+ using (_≟_; _≤?_; _<?_)
diff --git a/src/Data/Fin/Base.agda b/src/Data/Fin/Base.agda
new file mode 100644
index 0000000..a93915e
--- /dev/null
+++ b/src/Data/Fin/Base.agda
@@ -0,0 +1,239 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Finite sets
+------------------------------------------------------------------------
+
+-- Note that elements of Fin n can be seen as natural numbers in the
+-- set {m | m < n}. The notation "m" in comments below refers to this
+-- natural number view.
+
+module Data.Fin.Base where
+
+open import Data.Empty using (⊥-elim)
+open import Data.Nat as ℕ
+ using (ℕ; zero; suc; z≤n; s≤s)
+open import Function using (_∘_; _on_)
+open import Level using () renaming (zero to ℓ₀)
+open import Relation.Nullary using (yes; no)
+open import Relation.Nullary.Decidable using (True; toWitness)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality
+ using (_≡_; _≢_; refl; cong)
+
+------------------------------------------------------------------------
+-- Types
+
+-- Fin n is a type with n elements.
+
+data Fin : ℕ → Set where
+ zero : {n : ℕ} → Fin (suc n)
+ suc : {n : ℕ} (i : Fin n) → Fin (suc n)
+
+-- A conversion: toℕ "n" = n.
+
+toℕ : ∀ {n} → Fin n → ℕ
+toℕ zero = 0
+toℕ (suc i) = suc (toℕ i)
+
+-- A Fin-indexed variant of Fin.
+
+Fin′ : ∀ {n} → Fin n → Set
+Fin′ i = Fin (toℕ i)
+
+------------------------------------------------------------------------
+-- Conversions
+
+-- toℕ is defined above.
+
+-- fromℕ n = "n".
+
+fromℕ : (n : ℕ) → Fin (suc n)
+fromℕ zero = zero
+fromℕ (suc n) = suc (fromℕ n)
+
+-- fromℕ≤ {m} _ = "m".
+
+fromℕ≤ : ∀ {m n} → m ℕ.< n → Fin n
+fromℕ≤ (s≤s z≤n) = zero
+fromℕ≤ (s≤s (s≤s m≤n)) = suc (fromℕ≤ (s≤s m≤n))
+
+-- fromℕ≤″ m _ = "m".
+
+fromℕ≤″ : ∀ m {n} → m ℕ.<″ n → Fin n
+fromℕ≤″ zero (ℕ.less-than-or-equal refl) = zero
+fromℕ≤″ (suc m) (ℕ.less-than-or-equal refl) =
+ suc (fromℕ≤″ m (ℕ.less-than-or-equal refl))
+
+-- # m = "m".
+
+infix 10 #_
+
+#_ : ∀ m {n} {m<n : True (suc m ℕ.≤? n)} → Fin n
+#_ _ {m<n = m<n} = fromℕ≤ (toWitness m<n)
+
+-- raise m "n" = "m + n".
+
+raise : ∀ {m} n → Fin m → Fin (n ℕ.+ m)
+raise zero i = i
+raise (suc n) i = suc (raise n i)
+
+-- reduce≥ "m + n" _ = "n".
+
+reduce≥ : ∀ {m n} (i : Fin (m ℕ.+ n)) (i≥m : toℕ i ℕ.≥ m) → Fin n
+reduce≥ {zero} i i≥m = i
+reduce≥ {suc m} zero ()
+reduce≥ {suc m} (suc i) (s≤s i≥m) = reduce≥ i i≥m
+
+-- inject⋆ m "n" = "n".
+
+inject : ∀ {n} {i : Fin n} → Fin′ i → Fin n
+inject {i = zero} ()
+inject {i = suc i} zero = zero
+inject {i = suc i} (suc j) = suc (inject j)
+
+inject! : ∀ {n} {i : Fin (suc n)} → Fin′ i → Fin n
+inject! {n = zero} {i = suc ()} _
+inject! {i = zero} ()
+inject! {n = suc _} {i = suc _} zero = zero
+inject! {n = suc _} {i = suc _} (suc j) = suc (inject! j)
+
+inject+ : ∀ {m} n → Fin m → Fin (m ℕ.+ n)
+inject+ n zero = zero
+inject+ n (suc i) = suc (inject+ n i)
+
+inject₁ : ∀ {m} → Fin m → Fin (suc m)
+inject₁ zero = zero
+inject₁ (suc i) = suc (inject₁ i)
+
+inject≤ : ∀ {m n} → Fin m → m ℕ.≤ n → Fin n
+inject≤ zero (s≤s le) = zero
+inject≤ (suc i) (s≤s le) = suc (inject≤ i le)
+
+-- A strengthening injection into the minimal Fin fibre.
+strengthen : ∀ {n} (i : Fin n) → Fin′ (suc i)
+strengthen zero = zero
+strengthen (suc i) = suc (strengthen i)
+
+------------------------------------------------------------------------
+-- Operations
+
+-- Folds.
+
+fold : ∀ {t} (T : ℕ → Set t) {m} →
+ (∀ {n} → T n → T (suc n)) →
+ (∀ {n} → T (suc n)) →
+ Fin m → T m
+fold T f x zero = x
+fold T f x (suc i) = f (fold T f x i)
+
+fold′ : ∀ {n t} (T : Fin (suc n) → Set t) →
+ (∀ i → T (inject₁ i) → T (suc i)) →
+ T zero →
+ ∀ i → T i
+fold′ T f x zero = x
+fold′ {n = zero} T f x (suc ())
+fold′ {n = suc n} T f x (suc i) =
+ f i (fold′ (T ∘ inject₁) (f ∘ inject₁) x i)
+
+-- Lifts functions.
+
+lift : ∀ {m n} k → (Fin m → Fin n) → Fin (k ℕ.+ m) → Fin (k ℕ.+ n)
+lift zero f i = f i
+lift (suc k) f zero = zero
+lift (suc k) f (suc i) = suc (lift k f i)
+
+-- "m" + "n" = "m + n".
+
+infixl 6 _+_
+
+_+_ : ∀ {m n} (i : Fin m) (j : Fin n) → Fin (toℕ i ℕ.+ n)
+zero + j = j
+suc i + j = suc (i + j)
+
+-- "m" - "n" = "m ∸ n".
+
+infixl 6 _-_
+
+_-_ : ∀ {m} (i : Fin m) (j : Fin′ (suc i)) → Fin (m ℕ.∸ toℕ j)
+i - zero = i
+zero - suc ()
+suc i - suc j = i - j
+
+-- m ℕ- "n" = "m ∸ n".
+
+infixl 6 _ℕ-_
+
+_ℕ-_ : (n : ℕ) (j : Fin (suc n)) → Fin (suc n ℕ.∸ toℕ j)
+n ℕ- zero = fromℕ n
+zero ℕ- suc ()
+suc n ℕ- suc i = n ℕ- i
+
+-- m ℕ-ℕ "n" = m ∸ n.
+
+infixl 6 _ℕ-ℕ_
+
+_ℕ-ℕ_ : (n : ℕ) → Fin (suc n) → ℕ
+n ℕ-ℕ zero = n
+zero ℕ-ℕ suc ()
+suc n ℕ-ℕ suc i = n ℕ-ℕ i
+
+-- pred "n" = "pred n".
+
+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
+
+infix 4 _≤_ _<_
+
+_≤_ : ∀ {n} → Rel (Fin n) ℓ₀
+_≤_ = ℕ._≤_ on toℕ
+
+_<_ : ∀ {n} → Rel (Fin n) ℓ₀
+_<_ = ℕ._<_ on toℕ
+
+data _≺_ : ℕ → ℕ → Set where
+ _≻toℕ_ : ∀ n (i : Fin n) → toℕ i ≺ n
+
+------------------------------------------------------------------------
+-- An ordering view.
+
+data Ordering {n : ℕ} : Fin n → Fin n → Set where
+ less : ∀ greatest (least : Fin′ greatest) →
+ Ordering (inject least) greatest
+ equal : ∀ i → Ordering i i
+ greater : ∀ greatest (least : Fin′ greatest) →
+ Ordering greatest (inject least)
+
+compare : ∀ {n} (i j : Fin n) → Ordering i j
+compare zero zero = equal zero
+compare zero (suc j) = less (suc j) zero
+compare (suc i) zero = greater (suc i) zero
+compare (suc i) (suc j) with compare i j
+compare (suc .(inject least)) (suc .greatest) | less greatest least =
+ less (suc greatest) (suc least)
+compare (suc .greatest) (suc .(inject least)) | greater greatest least =
+ greater (suc greatest) (suc least)
+compare (suc .i) (suc .i) | equal i =
+ equal (suc i)
diff --git a/src/Data/Fin/Dec.agda b/src/Data/Fin/Dec.agda
index 2a31707..7607a4e 100644
--- a/src/Data/Fin/Dec.agda
+++ b/src/Data/Fin/Dec.agda
@@ -2,174 +2,16 @@
-- The Agda standard library
--
-- Decision procedures for finite sets and subsets of finite sets
+--
+-- This module is DEPRECATED. Please use the Data.Fin.Properties
+-- and Data.Fin.Subset.Properties directly.
------------------------------------------------------------------------
module Data.Fin.Dec where
-open import Function
-import Data.Bool as Bool
-open import Data.Nat.Base hiding (_<_)
-open import Data.Vec hiding (_∈_)
-open import Data.Vec.Equality as VecEq
- using () renaming (module PropositionalEquality to PropVecEq)
-open import Data.Fin
-open import Data.Fin.Subset
-open import Data.Fin.Subset.Properties
-open import Data.Product as Prod
-open import Data.Empty
-open import Function
-import Function.Equivalence as Eq
-open import Relation.Binary as B
-import Relation.Binary.HeterogeneousEquality as H
-open import Relation.Nullary
-import Relation.Nullary.Decidable as Dec
-open import Relation.Unary as U using (Pred)
-
-infix 4 _∈?_
-
-_∈?_ : ∀ {n} x (p : Subset n) → Dec (x ∈ p)
-zero ∈? inside ∷ p = yes here
-zero ∈? outside ∷ p = no λ()
-suc n ∈? s ∷ p with n ∈? p
-... | yes n∈p = yes (there n∈p)
-... | no n∉p = no (n∉p ∘ drop-there)
-
-private
-
- restrictP : ∀ {p n} → (Fin (suc n) → Set p) → (Fin n → Set p)
- restrictP P f = P (suc f)
-
- restrict : ∀ {p n} {P : Fin (suc n) → Set p} →
- U.Decidable P → U.Decidable (restrictP P)
- restrict dec f = dec (suc f)
-
-any? : ∀ {n} {P : Fin n → Set} →
- U.Decidable P → Dec (∃ P)
-any? {zero} dec = no λ { (() , _) }
-any? {suc n} {P} dec with dec zero | any? (restrict dec)
-... | yes p | _ = yes (_ , p)
-... | _ | yes (_ , p') = yes (_ , p')
-... | no ¬p | no ¬p' = no helper
- where
- helper : ∄ P
- helper (zero , p) = ¬p p
- helper (suc f , p') = ¬p' (_ , p')
-
-nonempty? : ∀ {n} (p : Subset n) → Dec (Nonempty p)
-nonempty? p = any? (λ x → x ∈? p)
-
-private
-
- restrict∈ : ∀ {p q n}
- (P : Fin (suc n) → Set p) {Q : Fin (suc n) → Set q} →
- (∀ {f} → Q f → Dec (P f)) →
- (∀ {f} → restrictP Q f → Dec (restrictP P f))
- restrict∈ _ dec {f} Qf = dec {suc f} Qf
-
-decFinSubset : ∀ {p q n} {P : Fin n → Set p} {Q : Fin n → Set q} →
- U.Decidable Q →
- (∀ {f} → Q f → Dec (P f)) →
- Dec (∀ {f} → Q f → P f)
-decFinSubset {n = zero} _ _ = yes λ{}
-decFinSubset {n = suc n} {P} {Q} decQ decP = helper
- where
- helper : Dec (∀ {f} → Q f → P f)
- helper with decFinSubset (restrict decQ) (restrict∈ P decP)
- helper | no ¬q⟶p = no (λ q⟶p → ¬q⟶p (λ {f} q → q⟶p {suc f} q))
- helper | yes q⟶p with decQ zero
- helper | yes q⟶p | yes q₀ with decP q₀
- helper | yes q⟶p | yes q₀ | no ¬p₀ = no (λ q⟶p → ¬p₀ (q⟶p {zero} q₀))
- helper | yes q⟶p | yes q₀ | yes p₀ = yes (λ {_} → hlpr _)
- where
- hlpr : ∀ f → Q f → P f
- hlpr zero _ = p₀
- hlpr (suc f) qf = q⟶p qf
- helper | yes q⟶p | no ¬q₀ = yes (λ {_} → hlpr _)
- where
- hlpr : ∀ f → Q f → P f
- hlpr zero q₀ = ⊥-elim (¬q₀ q₀)
- hlpr (suc f) qf = q⟶p qf
-
-all∈? : ∀ {n p} {P : Fin n → Set p} {q} →
- (∀ {f} → f ∈ q → Dec (P f)) →
- Dec (∀ {f} → f ∈ q → P f)
-all∈? {q = q} dec = decFinSubset (λ f → f ∈? q) dec
-
-all? : ∀ {n p} {P : Fin n → Set p} →
- U.Decidable P → Dec (∀ f → P f)
-all? dec with all∈? {q = ⊤} (λ {f} _ → dec f)
-... | yes ∀p = yes (λ f → ∀p ∈⊤)
-... | no ¬∀p = no (λ ∀p → ¬∀p (λ {f} _ → ∀p f))
-
-decLift : ∀ {n} {P : Fin n → Set} →
- U.Decidable P → U.Decidable (Lift P)
-decLift dec p = all∈? (λ {x} _ → dec x)
-
-private
-
- restrictSP : ∀ {n} → Side → (Subset (suc n) → Set) → (Subset n → Set)
- restrictSP s P p = P (s ∷ p)
-
- restrictS : ∀ {n} {P : Subset (suc n) → Set} →
- (s : Side) → U.Decidable P → U.Decidable (restrictSP s P)
- restrictS s dec p = dec (s ∷ p)
-
-anySubset? : ∀ {n} {P : Subset n → Set} →
- U.Decidable P → Dec (∃ P)
-anySubset? {zero} {P} dec with dec []
-... | yes P[] = yes (_ , P[])
-... | no ¬P[] = no helper
- where
- helper : ∄ P
- helper ([] , P[]) = ¬P[] P[]
-anySubset? {suc n} {P} dec with anySubset? (restrictS inside dec)
- | anySubset? (restrictS outside dec)
-... | yes (_ , Pp) | _ = yes (_ , Pp)
-... | _ | yes (_ , Pp) = yes (_ , Pp)
-... | no ¬Pp | no ¬Pp' = no helper
- where
- helper : ∄ P
- helper (inside ∷ p , Pp) = ¬Pp (_ , Pp)
- helper (outside ∷ p , Pp') = ¬Pp' (_ , Pp')
-
--- If a decidable predicate P over a finite set is sometimes false,
--- then we can find the smallest value for which this is the case.
-
-¬∀⟶∃¬-smallest :
- ∀ n {p} (P : Fin n → Set p) → U.Decidable P →
- ¬ (∀ i → P i) → ∃ λ i → ¬ P i × ((j : Fin′ i) → P (inject j))
-¬∀⟶∃¬-smallest zero P dec ¬∀iPi = ⊥-elim (¬∀iPi (λ()))
-¬∀⟶∃¬-smallest (suc n) P dec ¬∀iPi with dec zero
-¬∀⟶∃¬-smallest (suc n) P dec ¬∀iPi | no ¬P0 = (zero , ¬P0 , λ ())
-¬∀⟶∃¬-smallest (suc n) P dec ¬∀iPi | yes P0 =
- Prod.map suc (Prod.map id extend′) $
- ¬∀⟶∃¬-smallest n (λ n → P (suc n)) (dec ∘ suc) (¬∀iPi ∘ extend)
- where
- extend : (∀ i → P (suc i)) → (∀ i → P i)
- extend ∀iP[1+i] zero = P0
- extend ∀iP[1+i] (suc i) = ∀iP[1+i] i
-
- extend′ : ∀ {i : Fin n} →
- ((j : Fin′ i) → P (suc (inject j))) →
- ((j : Fin′ (suc i)) → P (inject j))
- 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).
-
-infix 4 _⊆?_
+open import Data.Fin.Properties public
+ using (decFinSubset; any?; all?; ¬∀⟶∃¬-smallest; ¬∀⟶∃¬)
-_⊆?_ : ∀ {n} → B.Decidable (_⊆_ {n = n})
-p₁ ⊆? p₂ =
- Dec.map (Eq.sym NaturalPoset.orders-equivalent) $
- Dec.map′ PropVecEq.to-≡ PropVecEq.from-≡ $
- VecEq.DecidableEquality._≟_ Bool.decSetoid p₁ (p₁ ∩ p₂)
+open import Data.Fin.Subset.Properties public
+ using (_∈?_; _⊆?_; nonempty?; anySubset?)
+ renaming (Lift? to decLift)
diff --git a/src/Data/Fin/Literals.agda b/src/Data/Fin/Literals.agda
new file mode 100644
index 0000000..1ecf01e
--- /dev/null
+++ b/src/Data/Fin/Literals.agda
@@ -0,0 +1,18 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Fin Literals
+------------------------------------------------------------------------
+
+module Data.Fin.Literals where
+
+open import Agda.Builtin.FromNat
+open import Data.Nat using (suc; _≤?_)
+open import Data.Fin using (Fin ; #_)
+open import Relation.Nullary.Decidable using (True)
+
+number : ∀ n → Number (Fin n)
+number n = record
+ { Constraint = λ m → True (suc m ≤? n)
+ ; fromNat = λ m {{pr}} → (# m) {n} {pr}
+ }
diff --git a/src/Data/Fin/Permutation.agda b/src/Data/Fin/Permutation.agda
new file mode 100644
index 0000000..749d15c
--- /dev/null
+++ b/src/Data/Fin/Permutation.agda
@@ -0,0 +1,183 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Bijections on finite sets (i.e. permutations).
+------------------------------------------------------------------------
+
+module Data.Fin.Permutation where
+
+open import Data.Empty using (⊥-elim)
+open import Data.Fin
+open import Data.Fin.Properties
+import Data.Fin.Permutation.Components as PC
+open import Data.Nat using (ℕ; suc; zero)
+open import Data.Product using (proj₂)
+open import Function.Inverse as Inverse using (_↔_; Inverse; _InverseOf_)
+open import Function.Equality using (_⟨$⟩_)
+open import Function using (_∘_)
+open import Relation.Nullary using (¬_; yes; no)
+open import Relation.Nullary.Negation using (contradiction)
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≢_; refl; trans; sym; →-to-⟶; cong; cong₂)
+open P.≡-Reasoning
+
+------------------------------------------------------------------------
+-- Types
+
+-- A bijection between finite sets of potentially different sizes.
+-- There only exist inhabitants of this type if in fact m = n, however
+-- it is often easier to prove the existence of a bijection without
+-- first proving that the sets have the same size. Indeed such a
+-- bijection is a useful way to prove that the sets are in fact the same
+-- size. See '↔-≡' below.
+
+Permutation : ℕ → ℕ → Set
+Permutation m n = Fin m ↔ Fin n
+
+Permutation′ : ℕ → Set
+Permutation′ n = Permutation n n
+
+------------------------------------------------------------------------
+-- Helper functions
+
+permutation : ∀ {m n} (f : Fin m → Fin n) (g : Fin n → Fin m) →
+ (→-to-⟶ g) InverseOf (→-to-⟶ f) → Permutation m n
+permutation f g inv = record
+ { to = →-to-⟶ f
+ ; from = →-to-⟶ g
+ ; inverse-of = inv
+ }
+
+_⟨$⟩ʳ_ : ∀ {m n} → Permutation m n → Fin m → Fin n
+_⟨$⟩ʳ_ = _⟨$⟩_ ∘ Inverse.to
+
+_⟨$⟩ˡ_ : ∀ {m n} → Permutation m n → Fin n → Fin m
+_⟨$⟩ˡ_ = _⟨$⟩_ ∘ Inverse.from
+
+inverseˡ : ∀ {m n} (π : Permutation m n) {i} → π ⟨$⟩ˡ (π ⟨$⟩ʳ i) ≡ i
+inverseˡ π = Inverse.left-inverse-of π _
+
+inverseʳ : ∀ {m n} (π : Permutation m n) {i} → π ⟨$⟩ʳ (π ⟨$⟩ˡ i) ≡ i
+inverseʳ π = Inverse.right-inverse-of π _
+
+------------------------------------------------------------------------
+-- Example permutations
+
+-- Identity
+
+id : ∀ {n} → Permutation′ n
+id = Inverse.id
+
+-- Transpose two indices
+
+transpose : ∀ {n} → Fin n → Fin n → Permutation′ n
+transpose i j = permutation (PC.transpose i j) (PC.transpose j i)
+ record
+ { left-inverse-of = λ _ → PC.transpose-inverse _ _
+ ; right-inverse-of = λ _ → PC.transpose-inverse _ _
+ }
+
+-- Reverse the order of indices
+
+reverse : ∀ {n} → Permutation′ n
+reverse = permutation PC.reverse PC.reverse
+ record
+ { left-inverse-of = PC.reverse-involutive
+ ; right-inverse-of = PC.reverse-involutive
+ }
+
+------------------------------------------------------------------------
+-- Operations
+
+-- Composition
+
+_∘ₚ_ : ∀ {m n o} → Permutation m n → Permutation n o → Permutation m o
+π₁ ∘ₚ π₂ = π₂ Inverse.∘ π₁
+
+-- Flip
+
+flip : ∀ {m n} → Permutation m n → Permutation n m
+flip = Inverse.sym
+
+-- Element removal
+--
+-- `remove k [0 ↦ i₀, …, k ↦ iₖ, …, n ↦ iₙ]` yields
+--
+-- [0 ↦ i₀, …, k-1 ↦ iₖ₋₁, k ↦ iₖ₊₁, k+1 ↦ iₖ₊₂, …, n-1 ↦ iₙ]
+
+remove : ∀ {m n} → Fin (suc m) →
+ Permutation (suc m) (suc n) → Permutation m n
+remove {m} {n} i π = permutation to from
+ record
+ { left-inverse-of = left-inverse-of
+ ; right-inverse-of = right-inverse-of
+ }
+ where
+ πʳ = π ⟨$⟩ʳ_
+ πˡ = π ⟨$⟩ˡ_
+
+ permute-≢ : ∀ {i j} → i ≢ j → πʳ i ≢ πʳ j
+ permute-≢ p = p ∘ (Inverse.injective π)
+
+ to-punchOut : ∀ {j : Fin m} → πʳ i ≢ πʳ (punchIn i j)
+ to-punchOut = permute-≢ (punchInᵢ≢i _ _ ∘ sym)
+
+ from-punchOut : ∀ {j : Fin n} → i ≢ πˡ (punchIn (πʳ i) j)
+ from-punchOut {j} p = punchInᵢ≢i (πʳ i) j (sym (begin
+ πʳ i ≡⟨ cong πʳ p ⟩
+ πʳ (πˡ (punchIn (πʳ i) j)) ≡⟨ inverseʳ π ⟩
+ punchIn (πʳ i) j ∎))
+
+ to : Fin m → Fin n
+ to j = punchOut (to-punchOut {j})
+
+ from : Fin n → Fin m
+ from j = punchOut {j = πˡ (punchIn (πʳ i) j)} from-punchOut
+
+ left-inverse-of : ∀ j → from (to j) ≡ j
+ left-inverse-of j = begin
+ from (to j) ≡⟨⟩
+ punchOut {i = i} {πˡ (punchIn (πʳ i) (punchOut to-punchOut))} _ ≡⟨ punchOut-cong′ i (cong πˡ (punchIn-punchOut {i = πʳ i} _)) ⟩
+ punchOut {i = i} {πˡ (πʳ (punchIn i j))} _ ≡⟨ punchOut-cong i (inverseˡ π) ⟩
+ punchOut {i = i} {punchIn i j} _ ≡⟨ punchOut-punchIn i ⟩
+ j ∎
+
+ right-inverse-of : ∀ j → to (from j) ≡ j
+ right-inverse-of j = begin
+ to (from j) ≡⟨⟩
+ punchOut {i = πʳ i} {πʳ (punchIn i (punchOut from-punchOut))} _ ≡⟨ punchOut-cong′ (πʳ i) (cong πʳ (punchIn-punchOut {i = i} _)) ⟩
+ punchOut {i = πʳ i} {πʳ (πˡ (punchIn (πʳ i) j))} _ ≡⟨ punchOut-cong (πʳ i) (inverseʳ π) ⟩
+ punchOut {i = πʳ i} {punchIn (πʳ i) j} _ ≡⟨ punchOut-punchIn (πʳ i) ⟩
+ j ∎
+
+------------------------------------------------------------------------
+-- Other properties
+
+module _ {m n} (π : Permutation (suc m) (suc n)) where
+ private
+ πʳ = π ⟨$⟩ʳ_
+ πˡ = π ⟨$⟩ˡ_
+
+ punchIn-permute : ∀ i j → πʳ (punchIn i j) ≡ punchIn (πʳ i) (remove i π ⟨$⟩ʳ j)
+ punchIn-permute i j = begin
+ πʳ (punchIn i j) ≡⟨ sym (punchIn-punchOut {i = πʳ i} _) ⟩
+ punchIn (πʳ i) (punchOut {i = πʳ i} {πʳ (punchIn i j)} _) ≡⟨⟩
+ punchIn (πʳ i) (remove i π ⟨$⟩ʳ j) ∎
+
+ punchIn-permute′ : ∀ i j → πʳ (punchIn (πˡ i) j) ≡ punchIn i (remove (πˡ i) π ⟨$⟩ʳ j)
+ punchIn-permute′ i j = begin
+ πʳ (punchIn (πˡ i) j) ≡⟨ punchIn-permute _ _ ⟩
+ punchIn (πʳ (πˡ i)) (remove (πˡ i) π ⟨$⟩ʳ j) ≡⟨ cong₂ punchIn (inverseʳ π) refl ⟩
+ punchIn i (remove (πˡ i) π ⟨$⟩ʳ j) ∎
+
+↔⇒≡ : ∀ {m n} → Permutation m n → m ≡ n
+↔⇒≡ {zero} {zero} π = refl
+↔⇒≡ {zero} {suc n} π = contradiction (π ⟨$⟩ˡ zero) ¬Fin0
+↔⇒≡ {suc m} {zero} π = contradiction (π ⟨$⟩ʳ zero) ¬Fin0
+↔⇒≡ {suc m} {suc n} π = cong suc (↔⇒≡ (remove zero π))
+
+fromPermutation : ∀ {m n} → Permutation m n → Permutation′ m
+fromPermutation π = P.subst (Permutation _) (sym (↔⇒≡ π)) π
+
+refute : ∀ {m n} → m ≢ n → ¬ Permutation m n
+refute m≢n π = contradiction (↔⇒≡ π) m≢n
diff --git a/src/Data/Fin/Permutation/Components.agda b/src/Data/Fin/Permutation/Components.agda
new file mode 100644
index 0000000..f0fc14e
--- /dev/null
+++ b/src/Data/Fin/Permutation/Components.agda
@@ -0,0 +1,75 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Component functions of permutations found in `Data.Fin.Permutation`
+------------------------------------------------------------------------
+
+module Data.Fin.Permutation.Components where
+
+open import Data.Fin
+open import Data.Fin.Properties
+open import Data.Nat as ℕ using (zero; suc; _∸_)
+import Data.Nat.Properties as ℕₚ
+open import Data.Product using (proj₂)
+open import Relation.Nullary using (yes; no)
+open import Relation.Binary.PropositionalEquality
+open import Algebra.FunctionProperties using (Involutive)
+open ≡-Reasoning
+
+--------------------------------------------------------------------------------
+-- Functions
+--------------------------------------------------------------------------------
+
+-- 'tranpose i j' swaps the places of 'i' and 'j'.
+
+transpose : ∀ {n} → Fin n → Fin n → Fin n → Fin n
+transpose i j k with k ≟ i
+... | yes _ = j
+... | no _ with k ≟ j
+... | yes _ = i
+... | no _ = k
+
+-- reverse i = n ∸ 1 ∸ i
+
+reverse : ∀ {n} → Fin n → Fin n
+reverse {zero} ()
+reverse {suc n} i = inject≤ (n ℕ- i) (ℕₚ.n∸m≤n (toℕ i) (suc n))
+
+--------------------------------------------------------------------------------
+-- Properties
+--------------------------------------------------------------------------------
+
+transpose-inverse : ∀ {n} (i j : Fin n) {k} →
+ transpose i j (transpose j i k) ≡ k
+transpose-inverse i j {k} with k ≟ j
+... | yes p rewrite ≡-≟-identity _≟_ {a = i} refl = sym p
+... | no ¬p with k ≟ i
+transpose-inverse i j {k} | no ¬p | yes q with j ≟ i
+... | yes r = trans r (sym q)
+... | no ¬r rewrite ≡-≟-identity _≟_ {a = j} refl = sym q
+transpose-inverse i j {k} | no ¬p | no ¬q
+ rewrite proj₂ (≢-≟-identity _≟_ ¬q)
+ | proj₂ (≢-≟-identity _≟_ ¬p) = refl
+
+reverse-prop : ∀ {n} → (i : Fin n) → toℕ (reverse i) ≡ n ∸ suc (toℕ i)
+reverse-prop {zero} ()
+reverse-prop {suc n} i = begin
+ toℕ (inject≤ (n ℕ- i) _) ≡⟨ toℕ-inject≤ _ _ ⟩
+ toℕ (n ℕ- i) ≡⟨ toℕ‿ℕ- n i ⟩
+ n ∸ toℕ i ∎
+
+reverse-involutive : ∀ {n} → Involutive _≡_ (reverse {n})
+reverse-involutive {zero} ()
+reverse-involutive {suc n} i = toℕ-injective (begin
+ toℕ (reverse (reverse i)) ≡⟨ reverse-prop (reverse i) ⟩
+ n ∸ (toℕ (reverse i)) ≡⟨ cong (n ∸_) (reverse-prop i) ⟩
+ n ∸ (n ∸ (toℕ i)) ≡⟨ ℕₚ.m∸[m∸n]≡n (ℕₚ.≤-pred (toℕ<n i)) ⟩
+ toℕ i ∎)
+
+reverse-suc : ∀ {n}{i : Fin n} → toℕ (reverse (suc i)) ≡ toℕ (reverse i)
+reverse-suc {n} {i} = begin
+ toℕ (reverse (suc i)) ≡⟨ reverse-prop (suc i) ⟩
+ suc n ∸ suc (toℕ (suc i)) ≡⟨⟩
+ n ∸ toℕ (suc i) ≡⟨⟩
+ n ∸ suc (toℕ i) ≡⟨ sym (reverse-prop i) ⟩
+ toℕ (reverse i) ∎
diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda
index 5732040..ef7944b 100644
--- a/src/Data/Fin/Properties.agda
+++ b/src/Data/Fin/Properties.agda
@@ -7,42 +7,53 @@
module Data.Fin.Properties where
-open import Algebra
+open import Algebra.FunctionProperties using (Involutive)
+open import Category.Applicative using (RawApplicative)
+open import Category.Functor using (RawFunctor)
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 _ℕ≤_
+open import Data.Fin.Base
+open import Data.Nat as ℕ using (ℕ; zero; suc; s≤s; z≤n; _∸_)
+ renaming
+ ( _≤_ to _ℕ≤_
; _<_ to _ℕ<_
- ; _+_ to _ℕ+_)
-import Data.Nat.Properties as N
-open import Data.Product
-open import Function
-open import Function.Equality as FunS using (_⟨$⟩_)
+ ; _+_ to _ℕ+_
+ )
+import Data.Nat.Properties as ℕₚ
+open import Data.Unit using (tt)
+open import Data.Product using (∃; ∃₂; ∄; _×_; _,_; map; proj₁)
+open import Function using (_∘_; id)
open import Function.Injection using (_↣_)
-open import Algebra.FunctionProperties
-open import Relation.Nullary
-import Relation.Nullary.Decidable as Dec
-open import Relation.Binary
+open import Relation.Binary as B hiding (Decidable)
open import Relation.Binary.PropositionalEquality as P
- using (_≡_; _≢_; refl; cong; subst)
-open import Category.Functor
-open import Category.Applicative
+ using (_≡_; _≢_; refl; sym; trans; cong; subst; module ≡-Reasoning)
+open import Relation.Nullary using (¬_)
+import Relation.Nullary.Decidable as Dec
+open import Relation.Nullary.Negation using (contradiction)
+open import Relation.Nullary using (Dec; yes; no; ¬_)
+open import Relation.Unary as U using (U; Pred; Decidable; _⊆_)
+open import Relation.Unary.Properties using (U?)
------------------------------------------------------------------------
--- Equality properties
+-- Fin
-infix 4 _≟_
+¬Fin0 : ¬ Fin 0
+¬Fin0 ()
+
+------------------------------------------------------------------------
+-- Properties of _≡_
suc-injective : ∀ {o} {m n : Fin o} → Fin.suc m ≡ suc n → m ≡ n
suc-injective refl = refl
-_≟_ : {n : ℕ} → Decidable {A = Fin n} _≡_
+infix 4 _≟_
+
+_≟_ : ∀ {n} → B.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)
+... | no x≢y = no (x≢y ∘ suc-injective)
preorder : ℕ → Preorder _ _ _
preorder n = P.preorder (Fin n)
@@ -64,19 +75,7 @@ decSetoid n = record
}
------------------------------------------------------------------------
--- Converting between Fin n and Nat
-
-to-from : ∀ n → toℕ (fromℕ n) ≡ n
-to-from zero = refl
-to-from (suc n) = cong suc (to-from n)
-
-from-to : ∀ {n} (i : Fin n) → fromℕ (toℕ i) ≡ strengthen i
-from-to zero = refl
-from-to (suc i) = cong suc (from-to i)
-
-toℕ-strengthen : ∀ {n} (i : Fin n) → toℕ (strengthen i) ≡ toℕ i
-toℕ-strengthen zero = refl
-toℕ-strengthen (suc i) = cong suc (toℕ-strengthen i)
+-- toℕ
toℕ-injective : ∀ {n} {i j : Fin n} → toℕ i ≡ toℕ j → i ≡ j
toℕ-injective {zero} {} {} _
@@ -84,23 +83,45 @@ toℕ-injective {suc n} {zero} {zero} eq = refl
toℕ-injective {suc n} {zero} {suc j} ()
toℕ-injective {suc n} {suc i} {zero} ()
toℕ-injective {suc n} {suc i} {suc j} eq =
- cong suc (toℕ-injective (cong N.pred eq))
+ cong suc (toℕ-injective (cong ℕ.pred eq))
+
+toℕ-strengthen : ∀ {n} (i : Fin n) → toℕ (strengthen i) ≡ toℕ i
+toℕ-strengthen zero = refl
+toℕ-strengthen (suc i) = cong suc (toℕ-strengthen i)
+
+toℕ-raise : ∀ {m} n (i : Fin m) → toℕ (raise n i) ≡ n ℕ+ toℕ i
+toℕ-raise zero i = refl
+toℕ-raise (suc n) i = cong suc (toℕ-raise n i)
-bounded : ∀ {n} (i : Fin n) → toℕ i ℕ< n
-bounded zero = s≤s z≤n
-bounded (suc i) = s≤s (bounded i)
+toℕ<n : ∀ {n} (i : Fin n) → toℕ i ℕ< n
+toℕ<n zero = s≤s z≤n
+toℕ<n (suc i) = s≤s (toℕ<n i)
-prop-toℕ-≤ : ∀ {n} (i : Fin n) → toℕ i ℕ≤ N.pred n
-prop-toℕ-≤ zero = z≤n
-prop-toℕ-≤ (suc {n = zero} ())
-prop-toℕ-≤ (suc {n = suc n} i) = s≤s (prop-toℕ-≤ i)
+toℕ≤pred[n] : ∀ {n} (i : Fin n) → toℕ i ℕ≤ ℕ.pred n
+toℕ≤pred[n] zero = z≤n
+toℕ≤pred[n] (suc {n = zero} ())
+toℕ≤pred[n] (suc {n = suc n} i) = s≤s (toℕ≤pred[n] i)
--- A simpler implementation of prop-toℕ-≤,
+-- A simpler implementation of toℕ≤pred[n],
-- however, with a different reduction behavior.
--- If no one needs the reduction behavior of prop-toℕ-≤,
--- it can be removed in favor of prop-toℕ-≤′.
-prop-toℕ-≤′ : ∀ {n} (i : Fin n) → toℕ i ℕ≤ N.pred n
-prop-toℕ-≤′ i = N.<⇒≤pred (bounded i)
+-- If no one needs the reduction behavior of toℕ≤pred[n],
+-- it can be removed in favor of toℕ≤pred[n]′.
+toℕ≤pred[n]′ : ∀ {n} (i : Fin n) → toℕ i ℕ≤ ℕ.pred n
+toℕ≤pred[n]′ i = ℕₚ.<⇒≤pred (toℕ<n i)
+
+------------------------------------------------------------------------
+-- fromℕ
+
+toℕ-fromℕ : ∀ n → toℕ (fromℕ n) ≡ n
+toℕ-fromℕ zero = refl
+toℕ-fromℕ (suc n) = cong suc (toℕ-fromℕ n)
+
+fromℕ-toℕ : ∀ {n} (i : Fin n) → fromℕ (toℕ i) ≡ strengthen i
+fromℕ-toℕ zero = refl
+fromℕ-toℕ (suc i) = cong suc (fromℕ-toℕ i)
+
+------------------------------------------------------------------------
+-- fromℕ≤
fromℕ≤-toℕ : ∀ {m} (i : Fin m) (i<m : toℕ i ℕ< m) → fromℕ≤ i<m ≡ i
fromℕ≤-toℕ zero (s≤s z≤n) = refl
@@ -111,38 +132,52 @@ 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 : ∀ 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ℕ≤≡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))
+fromℕ≤≡fromℕ≤″ : ∀ {m n} (m<n : m ℕ< n) (m<″n : m ℕ.<″ n) →
+ fromℕ≤ m<n ≡ fromℕ≤″ m m<″n
+fromℕ≤≡fromℕ≤″ (s≤s z≤n) (ℕ.less-than-or-equal refl) = refl
+fromℕ≤≡fromℕ≤″ (s≤s (s≤s m<n)) (ℕ.less-than-or-equal refl) =
+ cong suc (fromℕ≤≡fromℕ≤″ (s≤s m<n) (ℕ.less-than-or-equal refl))
-------------------------------------------------------------------------
--- Ordering properties
+toℕ-fromℕ≤″ : ∀ {m n} (m<n : m ℕ.<″ n) → toℕ (fromℕ≤″ m m<n) ≡ m
+toℕ-fromℕ≤″ {m} {n} m<n = begin
+ toℕ (fromℕ≤″ m m<n) ≡⟨ cong toℕ (sym (fromℕ≤≡fromℕ≤″ _ m<n)) ⟩
+ toℕ (fromℕ≤ _) ≡⟨ toℕ-fromℕ≤ (ℕₚ.≤″⇒≤ m<n) ⟩
+ m ∎
+ where open ≡-Reasoning
--- _≤_ ordering
+------------------------------------------------------------------------
+-- Properties of _≤_
+-- Relational properties
≤-reflexive : ∀ {n} → _≡_ ⇒ (_≤_ {n})
-≤-reflexive refl = N.≤-refl
+≤-reflexive refl = ℕₚ.≤-refl
≤-refl : ∀ {n} → Reflexive (_≤_ {n})
≤-refl = ≤-reflexive refl
≤-trans : ∀ {n} → Transitive (_≤_ {n})
-≤-trans = N.≤-trans
+≤-trans = ℕₚ.≤-trans
≤-antisym : ∀ {n} → Antisymmetric _≡_ (_≤_ {n})
-≤-antisym x≤y y≤x = toℕ-injective (N.≤-antisym x≤y y≤x)
+≤-antisym x≤y y≤x = toℕ-injective (ℕₚ.≤-antisym x≤y y≤x)
≤-total : ∀ {n} → Total (_≤_ {n})
-≤-total x y = N.≤-total (toℕ x) (toℕ y)
+≤-total x y = ℕₚ.≤-total (toℕ x) (toℕ y)
+
+infix 4 _≤?_ _<?_
+
+_≤?_ : ∀ {n} → B.Decidable (_≤_ {n})
+a ≤? b = toℕ a ℕ.≤? toℕ b
+
+_<?_ : ∀ {n} → B.Decidable (_<_ {n})
+m <? n = suc (toℕ m) ℕ.≤? toℕ n
≤-isPreorder : ∀ {n} → IsPreorder _≡_ (_≤_ {n})
≤-isPreorder = record
@@ -151,216 +186,247 @@ fromℕ≤≡fromℕ≤″ (s≤s (s≤s m<n)) (N.less-than-or-equal refl) =
; trans = ≤-trans
}
+≤-preorder : ℕ → Preorder _ _ _
+≤-preorder n = record
+ { isPreorder = ≤-isPreorder {n}
+ }
+
≤-isPartialOrder : ∀ {n} → IsPartialOrder _≡_ (_≤_ {n})
≤-isPartialOrder = record
{ isPreorder = ≤-isPreorder
; antisym = ≤-antisym
}
+≤-poset : ℕ → Poset _ _ _
+≤-poset n = record
+ { isPartialOrder = ≤-isPartialOrder {n}
+ }
+
≤-isTotalOrder : ∀ {n} → IsTotalOrder _≡_ (_≤_ {n})
≤-isTotalOrder = record
{ isPartialOrder = ≤-isPartialOrder
; total = ≤-total
}
--- _<_ ordering
+≤-totalOrder : ℕ → TotalOrder _ _ _
+≤-totalOrder n = record
+ { isTotalOrder = ≤-isTotalOrder {n}
+ }
+
+≤-isDecTotalOrder : ∀ {n} → IsDecTotalOrder _≡_ (_≤_ {n})
+≤-isDecTotalOrder = record
+ { isTotalOrder = ≤-isTotalOrder
+ ; _≟_ = _≟_
+ ; _≤?_ = _≤?_
+ }
+
+≤-decTotalOrder : ℕ → DecTotalOrder _ _ _
+≤-decTotalOrder n = record
+ { isDecTotalOrder = ≤-isDecTotalOrder {n}
+ }
+
+-- Other properties
+≤-irrelevance : ∀ {n} → Irrelevant (_≤_ {n})
+≤-irrelevance = ℕₚ.≤-irrelevance
+
+------------------------------------------------------------------------
+-- Properties of _<_
+
+-- Relational properties
+<-irrefl : ∀ {n} → Irreflexive _≡_ (_<_ {n})
+<-irrefl refl = ℕₚ.<-irrefl refl
+
+<-asym : ∀ {n} → Asymmetric (_<_ {n})
+<-asym = ℕₚ.<-asym
<-trans : ∀ {n} → Transitive (_<_ {n})
-<-trans = N.<-trans
+<-trans = ℕₚ.<-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< i<j i≢j j≮i = tri< (s≤s i<j) (i≢j ∘ suc-injective) (j≮i ∘ ℕₚ.≤-pred)
+... | tri> i≮j i≢j j<i = tri> (i≮j ∘ ℕₚ.≤-pred) (i≢j ∘ suc-injective) (s≤s j<i)
+... | tri≈ i≮j i≡j j≮i = tri≈ (i≮j ∘ ℕₚ.≤-pred) (cong suc i≡j) (j≮i ∘ ℕₚ.≤-pred)
+
+<-respˡ-≡ : ∀ {n} → (_<_ {n}) Respectsˡ _≡_
+<-respˡ-≡ refl x≤y = x≤y
+
+<-respʳ-≡ : ∀ {n} → (_<_ {n}) Respectsʳ _≡_
+<-respʳ-≡ refl x≤y = x≤y
-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)
+<-resp₂-≡ : ∀ {n} → (_<_ {n}) Respects₂ _≡_
+<-resp₂-≡ = <-respʳ-≡ , <-respˡ-≡
+
+<-isStrictPartialOrder : ∀ {n} → IsStrictPartialOrder _≡_ (_<_ {n})
+<-isStrictPartialOrder = record
+ { isEquivalence = P.isEquivalence
+ ; irrefl = <-irrefl
+ ; trans = <-trans
+ ; <-resp-≈ = <-resp₂-≡
+ }
-_<?_ : ∀ {n} → Decidable (_<_ {n})
-m <? n = suc (toℕ m) N.≤? toℕ n
+<-strictPartialOrder : ℕ → StrictPartialOrder _ _ _
+<-strictPartialOrder n = record
+ { isStrictPartialOrder = <-isStrictPartialOrder {n}
+ }
<-isStrictTotalOrder : ∀ {n} → IsStrictTotalOrder _≡_ (_<_ {n})
<-isStrictTotalOrder = record
{ isEquivalence = P.isEquivalence
; trans = <-trans
- ; compare = cmp
+ ; compare = <-cmp
}
-strictTotalOrder : ℕ → StrictTotalOrder _ _ _
-strictTotalOrder n = record
- { Carrier = Fin n
- ; _≈_ = _≡_
- ; _<_ = _<_
- ; isStrictTotalOrder = <-isStrictTotalOrder
+<-strictTotalOrder : ℕ → StrictTotalOrder _ _ _
+<-strictTotalOrder n = record
+ { isStrictTotalOrder = <-isStrictTotalOrder {n}
}
-------------------------------------------------------------------------
--- Injection properties
+-- Other properties
+<-irrelevance : ∀ {n} → Irrelevant (_<_ {n})
+<-irrelevance = ℕₚ.<-irrelevance
--- Lemma: n - i ≤ n.
-nℕ-ℕi≤n : ∀ n i → n ℕ-ℕ i ℕ≤ n
-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 ⟩
- n ≤⟨ N.n≤1+n n ⟩
- suc n ∎
- where open N.≤-Reasoning
+<⇒≢ : ∀ {n} {i j : Fin n} → i < j → i ≢ j
+<⇒≢ i<i refl = ℕₚ.n≮n _ i<i
+
+≤∧≢⇒< : ∀ {n} {i j : Fin n} → i ≤ j → i ≢ j → i < j
+≤∧≢⇒< {i = zero} {zero} _ 0≢0 = contradiction refl 0≢0
+≤∧≢⇒< {i = zero} {suc j} _ _ = s≤s z≤n
+≤∧≢⇒< {i = suc i} {zero} ()
+≤∧≢⇒< {i = suc i} {suc j} (s≤s i≤j) 1+i≢1+j =
+ s≤s (≤∧≢⇒< i≤j (1+i≢1+j ∘ (cong suc)))
+
+------------------------------------------------------------------------
+-- inject
-inject-lemma : ∀ {n} {i : Fin n} (j : Fin′ i) →
+toℕ-inject : ∀ {n} {i : Fin n} (j : Fin′ i) →
toℕ (inject j) ≡ toℕ j
-inject-lemma {i = zero} ()
-inject-lemma {i = suc i} zero = refl
-inject-lemma {i = suc i} (suc j) = cong suc (inject-lemma j)
+toℕ-inject {i = zero} ()
+toℕ-inject {i = suc i} zero = refl
+toℕ-inject {i = suc i} (suc j) = cong suc (toℕ-inject j)
+
+------------------------------------------------------------------------
+-- inject+
-inject+-lemma : ∀ {m} n (i : Fin m) → toℕ i ≡ toℕ (inject+ n i)
-inject+-lemma n zero = refl
-inject+-lemma n (suc i) = cong suc (inject+-lemma n i)
+toℕ-inject+ : ∀ {m} n (i : Fin m) → toℕ i ≡ toℕ (inject+ n i)
+toℕ-inject+ n zero = refl
+toℕ-inject+ n (suc i) = cong suc (toℕ-inject+ n i)
+
+------------------------------------------------------------------------
+-- inject₁
-inject₁-lemma : ∀ {m} (i : Fin m) → toℕ (inject₁ i) ≡ toℕ i
-inject₁-lemma zero = refl
-inject₁-lemma (suc i) = cong suc (inject₁-lemma i)
+inject₁-injective : ∀ {n} {i j : Fin n} → inject₁ i ≡ inject₁ j → i ≡ j
+inject₁-injective {i = zero} {zero} i≡j = refl
+inject₁-injective {i = zero} {suc j} ()
+inject₁-injective {i = suc i} {zero} ()
+inject₁-injective {i = suc i} {suc j} i≡j =
+ cong suc (inject₁-injective (suc-injective i≡j))
-inject≤-lemma : ∀ {m n} (i : Fin m) (le : m ℕ≤ n) →
+toℕ-inject₁ : ∀ {n} (i : Fin n) → toℕ (inject₁ i) ≡ toℕ i
+toℕ-inject₁ zero = refl
+toℕ-inject₁ (suc i) = cong suc (toℕ-inject₁ i)
+
+------------------------------------------------------------------------
+-- inject≤
+
+toℕ-inject≤ : ∀ {m n} (i : Fin m) (le : m ℕ≤ n) →
toℕ (inject≤ i le) ≡ toℕ i
-inject≤-lemma zero (N.s≤s le) = refl
-inject≤-lemma (suc i) (N.s≤s le) = cong suc (inject≤-lemma i le)
+toℕ-inject≤ zero (s≤s le) = refl
+toℕ-inject≤ (suc i) (s≤s le) = cong suc (toℕ-inject≤ i le)
--- Lemma: inject≤ i n≤n ≡ i.
inject≤-refl : ∀ {n} (i : Fin n) (n≤n : n ℕ≤ n) → inject≤ i n≤n ≡ i
inject≤-refl zero (s≤s _ ) = refl
inject≤-refl (suc i) (s≤s n≤n) = cong suc (inject≤-refl i n≤n)
-≺⇒<′ : _≺_ ⇒ N._<′_
-≺⇒<′ (n ≻toℕ i) = N.≤⇒≤′ (bounded i)
+------------------------------------------------------------------------
+-- _≺_
-<′⇒≺ : N._<′_ ⇒ _≺_
-<′⇒≺ {n} N.≤′-refl = subst (λ i → i ≺ suc n) (to-from n)
+≺⇒<′ : _≺_ ⇒ ℕ._<′_
+≺⇒<′ (n ≻toℕ i) = ℕₚ.≤⇒≤′ (toℕ<n i)
+
+<′⇒≺ : ℕ._<′_ ⇒ _≺_
+<′⇒≺ {n} ℕ.≤′-refl = subst (_≺ suc n) (toℕ-fromℕ n)
(suc n ≻toℕ fromℕ n)
-<′⇒≺ (N.≤′-step m≤′n) with <′⇒≺ m≤′n
-<′⇒≺ (N.≤′-step m≤′n) | n ≻toℕ i =
- subst (λ i → i ≺ suc n) (inject₁-lemma i) (suc n ≻toℕ (inject₁ i))
+<′⇒≺ (ℕ.≤′-step m≤′n) with <′⇒≺ m≤′n
+... | n ≻toℕ i = subst (_≺ suc n) (toℕ-inject₁ i) (suc n ≻toℕ _)
-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)
+------------------------------------------------------------------------
+-- pred
+
+<⇒≤pred : ∀ {n} {i j : Fin n} → j < i → j ≤ pred i
+<⇒≤pred {i = zero} {_} ()
+<⇒≤pred {i = suc i} {zero} j<i = z≤n
+<⇒≤pred {i = suc i} {suc j} (s≤s j<i) =
+ subst (_ ℕ≤_) (sym (toℕ-inject₁ i)) j<i
------------------------------------------------------------------------
--- Operations
+-- ℕ-
-infixl 6 _+′_
+toℕ‿ℕ- : ∀ n i → toℕ (n ℕ- i) ≡ n ∸ toℕ i
+toℕ‿ℕ- n zero = toℕ-fromℕ n
+toℕ‿ℕ- zero (suc ())
+toℕ‿ℕ- (suc n) (suc i) = toℕ‿ℕ- n i
-_+′_ : ∀ {m n} (i : Fin m) (j : Fin n) → Fin (N.pred m ℕ+ n)
-i +′ j = inject≤ (i + j) (N.+-mono-≤ (prop-toℕ-≤ i) N.≤-refl)
-
--- reverse {n} "i" = "n ∸ 1 ∸ i".
-
-reverse : ∀ {n} → Fin n → Fin n
-reverse {zero} ()
-reverse {suc n} i = inject≤ (n ℕ- i) (N.n∸m≤n (toℕ i) (suc n))
-
-reverse-prop : ∀ {n} → (i : Fin n) → toℕ (reverse i) ≡ n ∸ suc (toℕ i)
-reverse-prop {zero} ()
-reverse-prop {suc n} i = begin
- toℕ (inject≤ (n ℕ- i) _) ≡⟨ inject≤-lemma _ _ ⟩
- toℕ (n ℕ- i) ≡⟨ toℕ‿ℕ- n i ⟩
- n ∸ toℕ i ∎
- where
- open P.≡-Reasoning
-
- toℕ‿ℕ- : ∀ n i → toℕ (n ℕ- i) ≡ n ∸ toℕ i
- toℕ‿ℕ- n zero = to-from n
- toℕ‿ℕ- zero (suc ())
- toℕ‿ℕ- (suc n) (suc i) = toℕ‿ℕ- n i
-
-reverse-involutive : ∀ {n} → Involutive _≡_ reverse
-reverse-involutive {n} i = toℕ-injective (begin
- toℕ (reverse (reverse i)) ≡⟨ reverse-prop _ ⟩
- n ∸ suc (toℕ (reverse i)) ≡⟨ eq ⟩
- toℕ i ∎)
- where
- open P.≡-Reasoning
-
- lem₁ : ∀ m n → (m ℕ+ n) ∸ (m ℕ+ n ∸ m) ≡ m
- lem₁ m n = begin
- 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 ∎
-
- lem₂ : ∀ n → (i : Fin n) → n ∸ suc (n ∸ suc (toℕ i)) ≡ toℕ i
- lem₂ zero ()
- lem₂ (suc n) i = begin
- n ∸ (n ∸ toℕ i) ≡⟨ cong (λ ξ → ξ ∸ (ξ ∸ toℕ i)) i+j≡k ⟩
- (toℕ i ℕ+ j) ∸ (toℕ i ℕ+ j ∸ toℕ i) ≡⟨ lem₁ (toℕ i) j ⟩
- toℕ i ∎
- where
- decompose-n : ∃ λ j → n ≡ toℕ i ℕ+ j
- decompose-n = n ∸ toℕ i , P.sym (N.m+n∸m≡n (prop-toℕ-≤ i))
-
- j = proj₁ decompose-n
- i+j≡k = proj₂ decompose-n
-
- eq : n ∸ suc (toℕ (reverse i)) ≡ toℕ i
- eq = begin
- n ∸ suc (toℕ (reverse i)) ≡⟨ cong (λ ξ → n ∸ suc ξ) (reverse-prop i) ⟩
- n ∸ suc (n ∸ suc (toℕ i)) ≡⟨ lem₂ n i ⟩
- toℕ i ∎
-
--- Lemma: reverse {suc n} (suc i) ≡ reverse n i (in ℕ).
-
-reverse-suc : ∀{n}{i : Fin n} → toℕ (reverse (suc i)) ≡ toℕ (reverse i)
-reverse-suc {n}{i} = begin
- toℕ (reverse (suc i)) ≡⟨ reverse-prop (suc i) ⟩
- suc n ∸ suc (toℕ (suc i)) ≡⟨⟩
- n ∸ toℕ (suc i) ≡⟨⟩
- n ∸ suc (toℕ i) ≡⟨ P.sym (reverse-prop i) ⟩
- toℕ (reverse i) ∎
- where
- open P.≡-Reasoning
+------------------------------------------------------------------------
+-- ℕ-ℕ
--- If there is an injection from a type to a finite set, then the type
--- has decidable equality.
+nℕ-ℕi≤n : ∀ n i → n ℕ-ℕ i ℕ≤ n
+nℕ-ℕi≤n n zero = ℕₚ.≤-refl
+nℕ-ℕi≤n zero (suc ())
+nℕ-ℕi≤n (suc n) (suc i) = begin
+ n ℕ-ℕ i ≤⟨ nℕ-ℕi≤n n i ⟩
+ n ≤⟨ ℕₚ.n≤1+n n ⟩
+ suc n ∎
+ where open ℕₚ.≤-Reasoning
-eq? : ∀ {a n} {A : Set a} → A ↣ Fin n → Decidable {A = A} _≡_
-eq? inj = Dec.via-injection inj _≟_
+------------------------------------------------------------------------
+-- punchIn
--- Quantification over finite sets commutes with applicative functors.
+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))
-sequence : ∀ {F n} {P : Fin n → Set} → RawApplicative F →
- (∀ i → F (P i)) → F (∀ i → P i)
-sequence {F} RA = helper _ _
- where
- open RawApplicative RA
+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
- helper : ∀ n (P : Fin n → Set) → (∀ i → F (P i)) → F (∀ i → P i)
- helper zero P ∀iPi = pure (λ())
- helper (suc n) P ∀iPi =
- combine <$> ∀iPi zero ⊛ helper n (λ n → P (suc n)) (∀iPi ∘ suc)
- where
- combine : P zero → (∀ i → P (suc i)) → ∀ i → P i
- combine z s zero = z
- combine z s (suc i) = s i
+------------------------------------------------------------------------
+-- punchOut
-private
+-- A version of 'cong' for 'punchOut' in which the inequality argument can be
+-- changed out arbitrarily (reflecting the proof-irrelevance of that argument).
- -- Included just to show that sequence above has an inverse (under
- -- an equivalence relation with two equivalence classes, one with
- -- all inhabited sets and the other with all uninhabited sets).
+punchOut-cong : ∀ {n} (i : Fin (suc n)) {j k} {i≢j : i ≢ j} {i≢k : i ≢ k} → j ≡ k → punchOut i≢j ≡ punchOut i≢k
+punchOut-cong zero {zero} {i≢j = 0≢0} = contradiction refl 0≢0
+punchOut-cong zero {suc j} {zero} {i≢k = 0≢0} = contradiction refl 0≢0
+punchOut-cong zero {suc j} {suc k} = suc-injective
+punchOut-cong {zero} (suc ())
+punchOut-cong {suc n} (suc i) {zero} {zero} _ = refl
+punchOut-cong {suc n} (suc i) {zero} {suc k} ()
+punchOut-cong {suc n} (suc i) {suc j} {zero} ()
+punchOut-cong {suc n} (suc i) {suc j} {suc k} = cong suc ∘ punchOut-cong i ∘ suc-injective
- sequence⁻¹ : ∀ {F}{A} {P : A → Set} → RawFunctor F →
- F (∀ i → P i) → ∀ i → F (P i)
- sequence⁻¹ RF F∀iPi i = (λ f → f i) <$> F∀iPi
- where open RawFunctor RF
+-- An alternative to 'punchOut-cong' in the which the new inequality argument is
+-- specific. Useful for enabling the omission of that argument during equational
+-- reasoning.
-------------------------------------------------------------------------
+punchOut-cong′ : ∀ {n} (i : Fin (suc n)) {j k} {p : i ≢ j} (q : j ≡ k) → punchOut p ≡ punchOut (p ∘ sym ∘ trans q ∘ sym)
+punchOut-cong′ i q = punchOut-cong i q
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} {zero} {_} 0≢0 _ _ = contradiction refl 0≢0
+punchOut-injective {_} {zero} {_} {zero} _ 0≢0 _ = contradiction refl 0≢0
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
@@ -369,25 +435,197 @@ 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} {zero} 0≢0 = contradiction refl 0≢0
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
+punchOut-punchIn : ∀ {n} i {j : Fin n} → punchOut {i = i} {j = punchIn i j} (punchInᵢ≢i i j ∘ sym) ≡ j
+punchOut-punchIn zero {j} = refl
+punchOut-punchIn (suc i) {zero} = refl
+punchOut-punchIn (suc i) {suc j} = cong suc (begin
+ punchOut (punchInᵢ≢i i j ∘ suc-injective ∘ sym ∘ cong suc) ≡⟨ punchOut-cong i refl ⟩
+ punchOut (punchInᵢ≢i i j ∘ sym) ≡⟨ punchOut-punchIn i ⟩
+ j ∎)
+ where open ≡-Reasoning
+
+------------------------------------------------------------------------
+-- _+′_
+
+infixl 6 _+′_
+
+_+′_ : ∀ {m n} (i : Fin m) (j : Fin n) → Fin (ℕ.pred m ℕ+ n)
+i +′ j = inject≤ (i + j) (ℕₚ.+-mono-≤ (toℕ≤pred[n] i) ℕₚ.≤-refl)
+
+------------------------------------------------------------------------
+-- Quantification
+
+∀-cons : ∀ {n p} {P : Pred (Fin (suc n)) p} →
+ P zero → (∀ i → P (suc i)) → (∀ i → P i)
+∀-cons z s zero = z
+∀-cons z s (suc i) = s i
+
+decFinSubset : ∀ {n p q} {P : Pred (Fin n) p} {Q : Pred (Fin n) q} →
+ Decidable Q → (∀ {f} → Q f → Dec (P f)) → Dec (Q ⊆ P)
+decFinSubset {zero} {_} {_} _ _ = yes λ{}
+decFinSubset {suc n} {P = P} {Q} Q? P? with decFinSubset (Q? ∘ suc) P?
+... | no ¬q⟶p = no (λ q⟶p → ¬q⟶p (q⟶p))
+... | yes q⟶p with Q? zero
+... | no ¬q₀ = yes (∀-cons {P = Q U.⇒ P} (⊥-elim ∘ ¬q₀) (λ _ → q⟶p) _)
+... | yes q₀ with P? q₀
+... | no ¬p₀ = no (λ q⟶p → ¬p₀ (q⟶p q₀))
+... | yes p₀ = yes (∀-cons {P = Q U.⇒ P} (λ _ → p₀) (λ _ → q⟶p) _)
+
+any? : ∀ {n p} {P : Fin n → Set p} → Decidable P → Dec (∃ P)
+any? {zero} {_} P? = no λ { (() , _) }
+any? {suc n} {P} P? with P? zero | any? (P? ∘ suc)
+... | yes P₀ | _ = yes (_ , P₀)
+... | no _ | yes (_ , P₁₊ᵢ) = yes (_ , P₁₊ᵢ)
+... | no ¬P₀ | no ¬P₁₊ᵢ = no λ
+ { (zero , P₀) → ¬P₀ P₀
+ ; (suc f , P₁₊ᵢ) → ¬P₁₊ᵢ (_ , P₁₊ᵢ)
+ }
+
+all? : ∀ {n p} {P : Pred (Fin n) p} →
+ Decidable P → Dec (∀ f → P f)
+all? P? with decFinSubset U? (λ {f} _ → P? f)
+... | yes ∀p = yes (λ f → ∀p tt)
+... | no ¬∀p = no (λ ∀p → ¬∀p (λ _ → ∀p _))
+
+-- If a decidable predicate P over a finite set is sometimes false,
+-- then we can find the smallest value for which this is the case.
+
+¬∀⟶∃¬-smallest : ∀ n {p} (P : Pred (Fin n) p) → Decidable P →
+ ¬ (∀ i → P i) → ∃ λ i → ¬ P i × ((j : Fin′ i) → P (inject j))
+¬∀⟶∃¬-smallest zero P P? ¬∀P = contradiction (λ()) ¬∀P
+¬∀⟶∃¬-smallest (suc n) P P? ¬∀P with P? zero
+... | no ¬P₀ = (zero , ¬P₀ , λ ())
+... | yes P₀ = map suc (map id (∀-cons P₀))
+ (¬∀⟶∃¬-smallest n (P ∘ suc) (P? ∘ suc) (¬∀P ∘ (∀-cons P₀)))
+
+-- When P is a decidable predicate over a finite set the following
+-- lemma can be proved.
+
+¬∀⟶∃¬ : ∀ n {p} (P : Pred (Fin n) p) → Decidable P →
+ ¬ (∀ i → P i) → (∃ λ i → ¬ P i)
+¬∀⟶∃¬ n P P? ¬P = map id proj₁ (¬∀⟶∃¬-smallest n P P? ¬P)
+
+-- The pigeonhole principle.
+
+pigeonhole : ∀ {m n} → m ℕ.< n → (f : Fin n → Fin m) →
+ ∃₂ λ i j → i ≢ j × f i ≡ f j
+pigeonhole (s≤s z≤n) f = contradiction (f zero) λ()
+pigeonhole (s≤s (s≤s m≤n)) f with any? (λ k → f zero ≟ f (suc k))
+... | yes (j , f₀≡fⱼ) = zero , suc j , (λ()) , f₀≡fⱼ
+... | no f₀≢fₖ with pigeonhole (s≤s m≤n) (λ j → punchOut (f₀≢fₖ ∘ (j ,_ )))
+... | (i , j , i≢j , fᵢ≡fⱼ) =
+ suc i , suc j , i≢j ∘ suc-injective ,
+ punchOut-injective (f₀≢fₖ ∘ (i ,_)) _ fᵢ≡fⱼ
+
+------------------------------------------------------------------------
+-- Categorical
+
+module _ {f} {F : Set f → Set f} (RA : RawApplicative F) where
+
+ open RawApplicative RA
+
+ sequence : ∀ {n} {P : Pred (Fin n) f} →
+ (∀ i → F (P i)) → F (∀ i → P i)
+ sequence {zero} ∀iPi = pure λ()
+ sequence {suc n} ∀iPi = ∀-cons <$> ∀iPi zero ⊛ sequence (∀iPi ∘ suc)
+
+module _ {f} {F : Set f → Set f} (RF : RawFunctor F) where
+
+ open RawFunctor RF
+
+ sequence⁻¹ : ∀ {A : Set f} {P : Pred A f} →
+ F (∀ i → P i) → (∀ i → F (P i))
+ sequence⁻¹ F∀iPi i = (λ f → f i) <$> F∀iPi
+
+------------------------------------------------------------------------
+-- If there is an injection from a type to a finite set, then the type
+-- has decidable equality.
+
+module _ {a} {A : Set a} where
+
+ eq? : ∀ {n} → A ↣ Fin n → B.Decidable {A = A} _≡_
+ eq? inj = Dec.via-injection inj _≟_
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+cmp = <-cmp
+{-# WARNING_ON_USAGE cmp
+"Warning: cmp was deprecated in v0.15.
+Please use <-cmp instead."
+#-}
+strictTotalOrder = <-strictTotalOrder
+{-# WARNING_ON_USAGE strictTotalOrder
+"Warning: strictTotalOrder was deprecated in v0.15.
+Please use <-strictTotalOrder instead."
+#-}
+
+-- Version 0.16
+
+to-from = toℕ-fromℕ
+{-# WARNING_ON_USAGE to-from
+"Warning: to-from was deprecated in v0.16.
+Please use toℕ-fromℕ instead."
+#-}
+from-to = fromℕ-toℕ
+{-# WARNING_ON_USAGE from-to
+"Warning: from-to was deprecated in v0.16.
+Please use fromℕ-toℕ instead."
+#-}
+bounded = toℕ<n
+{-# WARNING_ON_USAGE bounded
+"Warning: bounded was deprecated in v0.16.
+Please use toℕ<n instead."
+#-}
+prop-toℕ-≤ = toℕ≤pred[n]
+{-# WARNING_ON_USAGE prop-toℕ-≤
+"Warning: prop-toℕ-≤ was deprecated in v0.16.
+Please use toℕ≤pred[n] instead."
+#-}
+prop-toℕ-≤′ = toℕ≤pred[n]′
+{-# WARNING_ON_USAGE prop-toℕ-≤′
+"Warning: prop-toℕ-≤′ was deprecated in v0.16.
+Please use toℕ≤pred[n]′ instead."
+#-}
+inject-lemma = toℕ-inject
+{-# WARNING_ON_USAGE inject-lemma
+"Warning: inject-lemma was deprecated in v0.16.
+Please use toℕ-inject instead."
+#-}
+inject+-lemma = toℕ-inject+
+{-# WARNING_ON_USAGE inject+-lemma
+"Warning: inject+-lemma was deprecated in v0.16.
+Please use toℕ-inject+ instead."
+#-}
+inject₁-lemma = toℕ-inject₁
+{-# WARNING_ON_USAGE inject₁-lemma
+"Warning: inject₁-lemma was deprecated in v0.16.
+Please use toℕ-inject₁ instead."
+#-}
+inject≤-lemma = toℕ-inject≤
+{-# WARNING_ON_USAGE inject≤-lemma
+"Warning: inject≤-lemma was deprecated in v0.16.
+Please use toℕ-inject≤ instead."
+#-}
+
+-- Version 0.17
+
+≤+≢⇒< = ≤∧≢⇒<
+{-# WARNING_ON_USAGE ≤+≢⇒<
+"Warning: ≤+≢⇒< was deprecated in v0.17.
+Please use ≤∧≢⇒< instead."
+#-}
diff --git a/src/Data/Fin/Subset.agda b/src/Data/Fin/Subset.agda
index 9b16298..2759151 100644
--- a/src/Data/Fin/Subset.agda
+++ b/src/Data/Fin/Subset.agda
@@ -7,15 +7,16 @@
module Data.Fin.Subset where
open import Algebra
+open import Algebra.FunctionProperties using (Op₁; Op₂)
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.Bool using (not; _∧_; _∨_; _≟_)
+open import Data.Fin using (Fin; zero; suc)
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 Data.Nat using (ℕ)
+open import Data.Product using (∃)
+open import Data.Vec hiding (foldr; foldl)
+import Data.Vec.Relation.Pointwise.Extensional as Pointwise
open import Relation.Nullary
------------------------------------------------------------------------
@@ -32,6 +33,22 @@ Subset : ℕ → Set
Subset = Vec Side
------------------------------------------------------------------------
+-- Special subsets
+
+-- The empty subset
+⊥ : ∀ {n} → Subset n
+⊥ = replicate outside
+
+-- The full subset
+⊤ : ∀ {n} → Subset n
+⊤ = replicate inside
+
+-- A singleton subset, containing just the given element.
+⁅_⁆ : ∀ {n} → Fin n → Subset n
+⁅ zero ⁆ = inside ∷ ⊥
+⁅ suc i ⁆ = outside ∷ ⁅ i ⁆
+
+------------------------------------------------------------------------
-- Membership and subset predicates
infix 4 _∈_ _∉_ _⊆_ _⊈_
@@ -43,54 +60,41 @@ _∉_ : ∀ {n} → Fin n → Subset n → Set
x ∉ p = ¬ (x ∈ p)
_⊆_ : ∀ {n} → Subset n → Subset n → Set
-p₁ ⊆ p₂ = ∀ {x} → x ∈ p₁ → x ∈ p₂
+p ⊆ q = ∀ {x} → x ∈ p → x ∈ q
_⊈_ : ∀ {n} → Subset n → Subset n → Set
-p₁ ⊈ p₂ = ¬ (p₁ ⊆ p₂)
+p ⊈ q = ¬ (p ⊆ q)
------------------------------------------------------------------------
-- Set operations
--- Pointwise lifting of the usual boolean algebra for booleans gives
--- us a boolean algebra for subsets.
---
--- The underlying equality of the returned boolean algebra is
--- propositional equality.
-
-booleanAlgebra : ℕ → BooleanAlgebra _ _
-booleanAlgebra n =
- BoolAlgProp.replace-equality
- (BAExpr.lift BoolProp.booleanAlgebra n)
- Pointwise.Pointwise-≡
-
-private
- open module BA {n} = BooleanAlgebra (booleanAlgebra n) public
- using
- ( ⊥ -- The empty subset.
- ; ⊤ -- The subset containing all elements.
- )
- renaming
- ( _∨_ to _∪_ -- Binary union.
- ; _∧_ to _∩_ -- Binary intersection.
- ; ¬_ to ∁ -- Complement.
- )
+infixr 7 _∩_
+infixr 6 _∪_
--- A singleton subset, containing just the given element.
+-- Complement
+∁ : ∀ {n} → Op₁ (Subset n)
+∁ p = map not p
-⁅_⁆ : ∀ {n} → Fin n → Subset n
-⁅ zero ⁆ = inside ∷ ⊥
-⁅ suc i ⁆ = outside ∷ ⁅ i ⁆
+-- Union
+_∩_ : ∀ {n} → Op₂ (Subset n)
+p ∩ q = zipWith _∧_ p q
--- N-ary union.
+-- Intersection
+_∪_ : ∀ {n} → Op₂ (Subset n)
+p ∪ q = zipWith _∨_ p q
+-- N-ary union
⋃ : ∀ {n} → List (Subset n) → Subset n
⋃ = foldr _∪_ ⊥
--- N-ary intersection.
-
+-- N-ary intersection
⋂ : ∀ {n} → List (Subset n) → Subset n
⋂ = foldr _∩_ ⊤
+-- Size
+∣_∣ : ∀ {n} → Subset n → ℕ
+∣ p ∣ = count (_≟ inside) p
+
------------------------------------------------------------------------
-- Properties
@@ -100,7 +104,5 @@ Nonempty p = ∃ λ f → f ∈ p
Empty : ∀ {n} (p : Subset n) → Set
Empty p = ¬ Nonempty p
--- Point-wise lifting of properties.
-
-Lift : ∀ {n} → (Fin n → Set) → (Subset n → Set)
+Lift : ∀ {n ℓ} → (Fin n → Set ℓ) → (Subset n → Set ℓ)
Lift P p = ∀ {x} → x ∈ p → P x
diff --git a/src/Data/Fin/Subset/Properties.agda b/src/Data/Fin/Subset/Properties.agda
index 5271457..68c0395 100644
--- a/src/Data/Fin/Subset/Properties.agda
+++ b/src/Data/Fin/Subset/Properties.agda
@@ -7,22 +7,30 @@
module Data.Fin.Subset.Properties where
open import Algebra
-import Algebra.Properties.BooleanAlgebra as BoolProp
-open import Data.Empty using (⊥-elim)
+import Algebra.FunctionProperties as AlgebraicProperties
+import Algebra.Structures as AlgebraicStructures
+import Algebra.Properties.Lattice as L
+import Algebra.Properties.DistributiveLattice as DL
+import Algebra.Properties.BooleanAlgebra as BA
+open import Data.Bool.Base using (_≟_)
+open import Data.Bool.Properties
open import Data.Fin using (Fin; suc; zero)
open import Data.Fin.Subset
-open import Data.Nat.Base using (ℕ)
-open import Data.Product as Product
-open import Data.Sum as Sum
-open import Data.Vec hiding (_∈_)
-open import Function
-open import Function.Equality using (_⟨$⟩_)
-open import Function.Equivalence
- using (_⇔_; equivalence; module Equivalence)
-open import Relation.Binary
+open import Data.Fin.Properties using (any?; decFinSubset)
+open import Data.Nat.Base using (ℕ; zero; suc; z≤n; s≤s; _≤_)
+open import Data.Nat.Properties using (≤-step)
+open import Data.Product as Product using (∃; ∄; _×_; _,_)
+open import Data.Sum as Sum using (_⊎_; inj₁; inj₂)
+open import Data.Vec
+open import Data.Vec.Properties
+open import Function using (_∘_; const; id; case_of_)
+open import Function.Equivalence using (_⇔_; equivalence)
+open import Relation.Binary as B hiding (Decidable)
open import Relation.Binary.PropositionalEquality
- using (_≡_; refl; cong; subst)
+ using (_≡_; refl; cong; cong₂; subst; isEquivalence)
open import Relation.Nullary.Negation using (contradiction)
+open import Relation.Nullary using (Dec; yes; no)
+open import Relation.Unary using (Pred; Decidable)
------------------------------------------------------------------------
-- Constructor mangling
@@ -31,13 +39,42 @@ drop-there : ∀ {s n x} {p : Subset n} → suc x ∈ s ∷ p → x ∈ p
drop-there (there x∈p) = x∈p
drop-∷-⊆ : ∀ {n s₁ s₂} {p₁ p₂ : Subset n} → s₁ ∷ p₁ ⊆ s₂ ∷ p₂ → p₁ ⊆ p₂
-drop-∷-⊆ p₁s₁⊆p₂s₂ x∈p₁ = drop-there $ p₁s₁⊆p₂s₂ (there x∈p₁)
+drop-∷-⊆ p₁s₁⊆p₂s₂ x∈p₁ = drop-there (p₁s₁⊆p₂s₂ (there x∈p₁))
+
+------------------------------------------------------------------------
+-- _∈_
+
+infix 4 _∈?_
+_∈?_ : ∀ {n} x (p : Subset n) → Dec (x ∈ p)
+zero ∈? inside ∷ p = yes here
+zero ∈? outside ∷ p = no λ()
+suc n ∈? s ∷ p with n ∈? p
+... | yes n∈p = yes (there n∈p)
+... | no n∉p = no (n∉p ∘ drop-there)
+
+------------------------------------------------------------------------
+-- Empty
drop-∷-Empty : ∀ {n s} {p : Subset n} → Empty (s ∷ p) → Empty p
drop-∷-Empty ¬∃∈ (x , x∈p) = ¬∃∈ (suc x , there x∈p)
+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 ¬∃∈))
+
+nonempty? : ∀ {n} → Decidable (Nonempty {n})
+nonempty? p = any? (_∈? p)
+
+------------------------------------------------------------------------
+-- ∣_∣
+
+∣p∣≤n : ∀ {n} (p : Subset n) → ∣ p ∣ ≤ n
+∣p∣≤n = count≤n (_≟ inside)
+
------------------------------------------------------------------------
--- Properties involving ⊥
+-- ⊥
∉⊥ : ∀ {n} {x : Fin n} → x ∉ ⊥
∉⊥ (there p) = ∉⊥ p
@@ -45,14 +82,12 @@ drop-∷-Empty ¬∃∈ (x , x∈p) = ¬∃∈ (suc x , there x∈p)
⊥⊆ : ∀ {n} {p : Subset n} → ⊥ ⊆ p
⊥⊆ x∈⊥ = contradiction x∈⊥ ∉⊥
-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 ¬∃∈))
+∣⊥∣≡0 : ∀ n → ∣ ⊥ {n = n} ∣ ≡ 0
+∣⊥∣≡0 zero = refl
+∣⊥∣≡0 (suc n) = ∣⊥∣≡0 n
------------------------------------------------------------------------
--- Properties involving ⊤
+-- ⊤
∈⊤ : ∀ {n} {x : Fin n} → x ∈ ⊤
∈⊤ {x = zero} = here
@@ -61,8 +96,12 @@ Empty-unique {_} {outside ∷ p} ¬∃∈ =
⊆⊤ : ∀ {n} {p : Subset n} → p ⊆ ⊤
⊆⊤ = const ∈⊤
+∣⊤∣≡n : ∀ n → ∣ ⊤ {n = n} ∣ ≡ n
+∣⊤∣≡n zero = refl
+∣⊤∣≡n (suc n) = cong suc (∣⊤∣≡n n)
+
------------------------------------------------------------------------
--- Properties involving ⁅_⁆
+-- ⁅_⁆
x∈⁅x⁆ : ∀ {n} (x : Fin n) → x ∈ ⁅ x ⁆
x∈⁅x⁆ zero = here
@@ -78,20 +117,373 @@ x∈⁅y⁆⇔x≡y {_} {x} {y} = equivalence
(x∈⁅y⁆⇒x≡y y)
(λ x≡y → subst (λ y → x ∈ ⁅ y ⁆) x≡y (x∈⁅x⁆ x))
+∣⁅x⁆∣≡1 : ∀ {n} (i : Fin n) → ∣ ⁅ i ⁆ ∣ ≡ 1
+∣⁅x⁆∣≡1 {suc n} zero = cong suc (∣⊥∣≡0 n)
+∣⁅x⁆∣≡1 {_} (suc i) = ∣⁅x⁆∣≡1 i
+
+------------------------------------------------------------------------
+-- _⊆_
+
+⊆-refl : ∀ {n} → Reflexive (_⊆_ {n})
+⊆-refl = id
+
+⊆-reflexive : ∀ {n} → _≡_ ⇒ (_⊆_ {n})
+⊆-reflexive refl = ⊆-refl
+
+⊆-trans : ∀ {n} → Transitive (_⊆_ {n})
+⊆-trans p⊆q q⊆r x∈p = q⊆r (p⊆q x∈p)
+
+⊆-antisym : ∀ {n} → Antisymmetric _≡_ (_⊆_ {n})
+⊆-antisym {x = []} {[]} p⊆q q⊆p = refl
+⊆-antisym {x = x ∷ xs} {y ∷ ys} p⊆q q⊆p with x | y
+... | inside | inside = cong₂ _∷_ refl (⊆-antisym (drop-∷-⊆ p⊆q) (drop-∷-⊆ q⊆p))
+... | inside | outside = contradiction (p⊆q here) λ()
+... | outside | inside = contradiction (q⊆p here) λ()
+... | outside | outside = cong₂ _∷_ refl (⊆-antisym (drop-∷-⊆ p⊆q) (drop-∷-⊆ q⊆p))
+
+⊆-min : ∀ {n} → Minimum (_⊆_ {n}) ⊥
+⊆-min [] ()
+⊆-min (x ∷ xs) (there v∈⊥) = there (⊆-min xs v∈⊥)
+
+⊆-max : ∀ {n} → Maximum (_⊆_ {n}) ⊤
+⊆-max [] ()
+⊆-max (inside ∷ xs) here = here
+⊆-max (x ∷ xs) (there v∈xs) = there (⊆-max xs v∈xs)
+
+infix 4 _⊆?_
+_⊆?_ : ∀ {n} → B.Decidable (_⊆_ {n = n})
+[] ⊆? [] = yes id
+outside ∷ p ⊆? y ∷ q with p ⊆? q
+... | yes p⊆q = yes λ { (there v∈p) → there (p⊆q v∈p)}
+... | no p⊈q = no (p⊈q ∘ drop-∷-⊆)
+inside ∷ p ⊆? outside ∷ q = no (λ p⊆q → case (p⊆q here) of λ())
+inside ∷ p ⊆? inside ∷ q with p ⊆? q
+... | yes p⊆q = yes λ { here → here ; (there v) → there (p⊆q v)}
+... | no p⊈q = no (p⊈q ∘ drop-∷-⊆)
+
+module _ (n : ℕ) where
+
+ ⊆-isPreorder : IsPreorder _≡_ (_⊆_ {n})
+ ⊆-isPreorder = record
+ { isEquivalence = isEquivalence
+ ; reflexive = ⊆-reflexive
+ ; trans = ⊆-trans
+ }
+
+ ⊆-preorder : Preorder _ _ _
+ ⊆-preorder = record
+ { isPreorder = ⊆-isPreorder
+ }
+
+ ⊆-isPartialOrder : IsPartialOrder _≡_ (_⊆_ {n})
+ ⊆-isPartialOrder = record
+ { isPreorder = ⊆-isPreorder
+ ; antisym = ⊆-antisym
+ }
+
+ ⊆-poset : Poset _ _ _
+ ⊆-poset = record
+ { isPartialOrder = ⊆-isPartialOrder
+ }
+
+p⊆q⇒∣p∣<∣q∣ : ∀ {n} {p q : Subset n} → p ⊆ q → ∣ p ∣ ≤ ∣ q ∣
+p⊆q⇒∣p∣<∣q∣ {p = []} {[]} p⊆q = z≤n
+p⊆q⇒∣p∣<∣q∣ {p = outside ∷ p} {outside ∷ q} p⊆q = p⊆q⇒∣p∣<∣q∣ (drop-∷-⊆ p⊆q)
+p⊆q⇒∣p∣<∣q∣ {p = outside ∷ p} {inside ∷ q} p⊆q = ≤-step (p⊆q⇒∣p∣<∣q∣ (drop-∷-⊆ p⊆q))
+p⊆q⇒∣p∣<∣q∣ {p = inside ∷ p} {outside ∷ q} p⊆q = contradiction (p⊆q here) λ()
+p⊆q⇒∣p∣<∣q∣ {p = inside ∷ p} {inside ∷ q} p⊆q = s≤s (p⊆q⇒∣p∣<∣q∣ (drop-∷-⊆ p⊆q))
+
------------------------------------------------------------------------
--- 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
- )
+ open AlgebraicProperties {A = Subset n} _≡_
+
+ ∩-assoc : Associative _∩_
+ ∩-assoc = zipWith-assoc ∧-assoc
+
+ ∩-comm : Commutative _∩_
+ ∩-comm = zipWith-comm ∧-comm
+
+ ∩-idem : Idempotent _∩_
+ ∩-idem = zipWith-idem ∧-idem
+
+ ∩-identityˡ : LeftIdentity ⊤ _∩_
+ ∩-identityˡ = zipWith-identityˡ ∧-identityˡ
+
+ ∩-identityʳ : RightIdentity ⊤ _∩_
+ ∩-identityʳ = zipWith-identityʳ ∧-identityʳ
+
+ ∩-identity : Identity ⊤ _∩_
+ ∩-identity = ∩-identityˡ , ∩-identityʳ
+
+ ∩-zeroˡ : LeftZero ⊥ _∩_
+ ∩-zeroˡ = zipWith-zeroˡ ∧-zeroˡ
+
+ ∩-zeroʳ : RightZero ⊥ _∩_
+ ∩-zeroʳ = zipWith-zeroʳ ∧-zeroʳ
+
+ ∩-zero : Zero ⊥ _∩_
+ ∩-zero = ∩-zeroˡ , ∩-zeroʳ
+
+ ∩-inverseˡ : LeftInverse ⊥ ∁ _∩_
+ ∩-inverseˡ = zipWith-inverseˡ ∧-inverseˡ
+
+ ∩-inverseʳ : RightInverse ⊥ ∁ _∩_
+ ∩-inverseʳ = zipWith-inverseʳ ∧-inverseʳ
+
+ ∩-inverse : Inverse ⊥ ∁ _∩_
+ ∩-inverse = ∩-inverseˡ , ∩-inverseʳ
+
+module _ (n : ℕ) where
+
+ open AlgebraicStructures {A = Subset n} _≡_
+
+ ∩-isSemigroup : IsSemigroup _∩_
+ ∩-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∩-assoc
+ ; ∙-cong = cong₂ _∩_
+ }
+
+ ∩-semigroup : Semigroup _ _
+ ∩-semigroup = record
+ { isSemigroup = ∩-isSemigroup
+ }
+
+ ∩-isMonoid : IsMonoid _∩_ ⊤
+ ∩-isMonoid = record
+ { isSemigroup = ∩-isSemigroup
+ ; identity = ∩-identity
+ }
+
+ ∩-monoid : Monoid _ _
+ ∩-monoid = record
+ { isMonoid = ∩-isMonoid
+ }
+
+ ∩-isCommutativeMonoid : IsCommutativeMonoid _∩_ ⊤
+ ∩-isCommutativeMonoid = record
+ { isSemigroup = ∩-isSemigroup
+ ; identityˡ = ∩-identityˡ
+ ; comm = ∩-comm
+ }
+
+ ∩-commutativeMonoid : CommutativeMonoid _ _
+ ∩-commutativeMonoid = record
+ { isCommutativeMonoid = ∩-isCommutativeMonoid
+ }
+
+ ∩-isIdempotentCommutativeMonoid : IsIdempotentCommutativeMonoid _∩_ ⊤
+ ∩-isIdempotentCommutativeMonoid = record
+ { isCommutativeMonoid = ∩-isCommutativeMonoid
+ ; idem = ∩-idem
+ }
+
+ ∩-idempotentCommutativeMonoid : IdempotentCommutativeMonoid _ _
+ ∩-idempotentCommutativeMonoid = record
+ { isIdempotentCommutativeMonoid = ∩-isIdempotentCommutativeMonoid
+ }
+
+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) (_ ∷ q) (there ∈p∩q) = there (p∩q⊆p p q ∈p∩q)
+p∩q⊆p (outside ∷ p) (_ ∷ 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⁺
+
+------------------------------------------------------------------------
-- _∪_
+module _ {n : ℕ} where
+
+ open AlgebraicProperties {A = Subset n} _≡_
+
+ ∪-assoc : Associative _∪_
+ ∪-assoc = zipWith-assoc ∨-assoc
+
+ ∪-comm : Commutative _∪_
+ ∪-comm = zipWith-comm ∨-comm
+
+ ∪-idem : Idempotent _∪_
+ ∪-idem = zipWith-idem ∨-idem
+
+ ∪-identityˡ : LeftIdentity ⊥ _∪_
+ ∪-identityˡ = zipWith-identityˡ ∨-identityˡ
+
+ ∪-identityʳ : RightIdentity ⊥ _∪_
+ ∪-identityʳ = zipWith-identityʳ ∨-identityʳ
+
+ ∪-identity : Identity ⊥ _∪_
+ ∪-identity = ∪-identityˡ , ∪-identityʳ
+
+ ∪-zeroˡ : LeftZero ⊤ _∪_
+ ∪-zeroˡ = zipWith-zeroˡ ∨-zeroˡ
+
+ ∪-zeroʳ : RightZero ⊤ _∪_
+ ∪-zeroʳ = zipWith-zeroʳ ∨-zeroʳ
+
+ ∪-zero : Zero ⊤ _∪_
+ ∪-zero = ∪-zeroˡ , ∪-zeroʳ
+
+ ∪-inverseˡ : LeftInverse ⊤ ∁ _∪_
+ ∪-inverseˡ = zipWith-inverseˡ ∨-inverseˡ
+
+ ∪-inverseʳ : RightInverse ⊤ ∁ _∪_
+ ∪-inverseʳ = zipWith-inverseʳ ∨-inverseʳ
+
+ ∪-inverse : Inverse ⊤ ∁ _∪_
+ ∪-inverse = ∪-inverseˡ , ∪-inverseʳ
+
+ ∪-distribˡ-∩ : _∪_ DistributesOverˡ _∩_
+ ∪-distribˡ-∩ = zipWith-distribˡ ∨-distribˡ-∧
+
+ ∪-distribʳ-∩ : _∪_ DistributesOverʳ _∩_
+ ∪-distribʳ-∩ = zipWith-distribʳ ∨-distribʳ-∧
+
+ ∪-distrib-∩ : _∪_ DistributesOver _∩_
+ ∪-distrib-∩ = ∪-distribˡ-∩ , ∪-distribʳ-∩
+
+ ∩-distribˡ-∪ : _∩_ DistributesOverˡ _∪_
+ ∩-distribˡ-∪ = zipWith-distribˡ ∧-distribˡ-∨
+
+ ∩-distribʳ-∪ : _∩_ DistributesOverʳ _∪_
+ ∩-distribʳ-∪ = zipWith-distribʳ ∧-distribʳ-∨
+
+ ∩-distrib-∪ : _∩_ DistributesOver _∪_
+ ∩-distrib-∪ = ∩-distribˡ-∪ , ∩-distribʳ-∪
+
+ ∪-abs-∩ : _∪_ Absorbs _∩_
+ ∪-abs-∩ = zipWith-absorbs ∨-abs-∧
+
+ ∩-abs-∪ : _∩_ Absorbs _∪_
+ ∩-abs-∪ = zipWith-absorbs ∧-abs-∨
+
+module _ (n : ℕ) where
+
+ open AlgebraicStructures {A = Subset n} _≡_
+
+ ∪-isSemigroup : IsSemigroup _∪_
+ ∪-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = ∪-assoc
+ ; ∙-cong = cong₂ _∪_
+ }
+
+ ∪-semigroup : Semigroup _ _
+ ∪-semigroup = record
+ { isSemigroup = ∪-isSemigroup
+ }
+
+ ∪-isMonoid : IsMonoid _∪_ ⊥
+ ∪-isMonoid = record
+ { isSemigroup = ∪-isSemigroup
+ ; identity = ∪-identity
+ }
+
+ ∪-monoid : Monoid _ _
+ ∪-monoid = record
+ { isMonoid = ∪-isMonoid
+ }
+
+ ∪-isCommutativeMonoid : IsCommutativeMonoid _∪_ ⊥
+ ∪-isCommutativeMonoid = record
+ { isSemigroup = ∪-isSemigroup
+ ; identityˡ = ∪-identityˡ
+ ; comm = ∪-comm
+ }
+
+ ∪-commutativeMonoid : CommutativeMonoid _ _
+ ∪-commutativeMonoid = record
+ { isCommutativeMonoid = ∪-isCommutativeMonoid
+ }
+
+ ∪-isIdempotentCommutativeMonoid : IsIdempotentCommutativeMonoid _∪_ ⊥
+ ∪-isIdempotentCommutativeMonoid = record
+ { isCommutativeMonoid = ∪-isCommutativeMonoid
+ ; idem = ∪-idem
+ }
+
+ ∪-idempotentCommutativeMonoid : IdempotentCommutativeMonoid _ _
+ ∪-idempotentCommutativeMonoid = record
+ { isIdempotentCommutativeMonoid = ∪-isIdempotentCommutativeMonoid
+ }
+
+ ∪-∩-isLattice : IsLattice _∪_ _∩_
+ ∪-∩-isLattice = record
+ { isEquivalence = isEquivalence
+ ; ∨-comm = ∪-comm
+ ; ∨-assoc = ∪-assoc
+ ; ∨-cong = cong₂ _∪_
+ ; ∧-comm = ∩-comm
+ ; ∧-assoc = ∩-assoc
+ ; ∧-cong = cong₂ _∩_
+ ; absorptive = ∪-abs-∩ , ∩-abs-∪
+ }
+
+ ∪-∩-lattice : Lattice _ _
+ ∪-∩-lattice = record
+ { isLattice = ∪-∩-isLattice
+ }
+
+ ∪-∩-isDistributiveLattice : IsDistributiveLattice _∪_ _∩_
+ ∪-∩-isDistributiveLattice = record
+ { isLattice = ∪-∩-isLattice
+ ; ∨-∧-distribʳ = ∪-distribʳ-∩
+ }
+
+ ∪-∩-distributiveLattice : DistributiveLattice _ _
+ ∪-∩-distributiveLattice = record
+ { isDistributiveLattice = ∪-∩-isDistributiveLattice
+ }
+
+ ∪-∩-isBooleanAlgebra : IsBooleanAlgebra _∪_ _∩_ ∁ ⊤ ⊥
+ ∪-∩-isBooleanAlgebra = record
+ { isDistributiveLattice = ∪-∩-isDistributiveLattice
+ ; ∨-complementʳ = ∪-inverseʳ
+ ; ∧-complementʳ = ∩-inverseʳ
+ ; ¬-cong = cong ∁
+ }
+
+ ∪-∩-booleanAlgebra : BooleanAlgebra _ _
+ ∪-∩-booleanAlgebra = record
+ { isBooleanAlgebra = ∪-∩-isBooleanAlgebra
+ }
+
+ ∩-∪-isLattice : IsLattice _∩_ _∪_
+ ∩-∪-isLattice = L.∧-∨-isLattice ∪-∩-lattice
+
+ ∩-∪-lattice : Lattice _ _
+ ∩-∪-lattice = L.∧-∨-lattice ∪-∩-lattice
+
+ ∩-∪-isDistributiveLattice : IsDistributiveLattice _∩_ _∪_
+ ∩-∪-isDistributiveLattice = DL.∧-∨-isDistributiveLattice ∪-∩-distributiveLattice
+
+ ∩-∪-distributiveLattice : DistributiveLattice _ _
+ ∩-∪-distributiveLattice = DL.∧-∨-distributiveLattice ∪-∩-distributiveLattice
+
+ ∩-∪-isBooleanAlgebra : IsBooleanAlgebra _∩_ _∪_ ∁ ⊥ ⊤
+ ∩-∪-isBooleanAlgebra = BA.∧-∨-isBooleanAlgebra ∪-∩-booleanAlgebra
+
+ ∩-∪-booleanAlgebra : BooleanAlgebra _ _
+ ∩-∪-booleanAlgebra = BA.∧-∨-booleanAlgebra ∪-∩-booleanAlgebra
+
p⊆p∪q : ∀ {n p} (q : Subset n) → p ⊆ p ∪ q
p⊆p∪q [] ()
p⊆p∪q (s ∷ q) here = here
@@ -114,77 +506,25 @@ 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))
+------------------------------------------------------------------------
+-- Lift
-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⁺
+Lift? : ∀ {n p} {P : Pred (Fin n) p} → Decidable P → Decidable (Lift P)
+Lift? P? p = decFinSubset (_∈? p) (λ {x} _ → P? x)
------------------------------------------------------------------------
--- _⊆_ is a partial order
-
--- The "natural poset" associated with the boolean algebra.
-
-module NaturalPoset where
- private
- open module BA {n} = BoolProp (booleanAlgebra n) public
- using (poset)
- open module Po {n} = Poset (poset {n = n}) public hiding (refl)
-
- -- _⊆_ is equivalent to the natural lattice order.
-
- orders-equivalent : ∀ {n} {p₁ p₂ : Subset n} → p₁ ⊆ p₂ ⇔ p₁ ≤ p₂
- orders-equivalent = equivalence (to _ _) (from _ _)
- where
- to : ∀ {n} (p₁ p₂ : Subset n) → p₁ ⊆ p₂ → p₁ ≤ p₂
- to [] [] p₁⊆p₂ = refl
- to (inside ∷ p₁) (_ ∷ p₂) p₁⊆p₂ with p₁⊆p₂ here
- 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 cong head p₁≤p₂ = here
- from (_ ∷ p₁) (_ ∷ p₂) p₁≤p₂ (there xs[i]=x) =
- there (from p₁ p₂ (cong tail p₁≤p₂) xs[i]=x)
-
--- _⊆_ is a partial order.
-
-poset : ℕ → Poset _ _ _
-poset n = record
- { Carrier = Subset n
- ; _≈_ = _≡_
- ; _≤_ = _⊆_
- ; isPartialOrder = record
- { isPreorder = record
- { isEquivalence = isEquivalence
- ; reflexive = λ i≡j → from ⟨$⟩ reflexive i≡j
- ; trans = λ x⊆y y⊆z → from ⟨$⟩ trans (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆z)
- }
- ; antisym = λ x⊆y y⊆x → antisym (to ⟨$⟩ x⊆y) (to ⟨$⟩ y⊆x)
- }
+-- Other
+
+anySubset? : ∀ {n} {P : Subset n → Set} → Decidable P → Dec (∃ P)
+anySubset? {zero} P? with P? []
+... | yes P[] = yes (_ , P[])
+... | no ¬P[] = no (λ {([] , P[]) → ¬P[] P[]})
+anySubset? {suc n} P? with anySubset? (P? ∘ (inside ∷_))
+... | yes (_ , Pp) = yes (_ , Pp)
+... | no ¬Pp with anySubset? (P? ∘ (outside ∷_))
+... | yes (_ , Pp) = yes (_ , Pp)
+... | no ¬Pp' = no λ
+ { (inside ∷ p , Pp) → ¬Pp (_ , Pp)
+ ; (outside ∷ p , Pp') → ¬Pp' (_ , Pp')
}
- where
- open NaturalPoset
- open module E {p₁ p₂} =
- Equivalence (orders-equivalent {n = n} {p₁ = p₁} {p₂ = p₂})
+
diff --git a/src/Data/Fin/Substitution.agda b/src/Data/Fin/Substitution.agda
index 20a9c4a..ddc14b1 100644
--- a/src/Data/Fin/Substitution.agda
+++ b/src/Data/Fin/Substitution.agda
@@ -13,11 +13,15 @@
module Data.Fin.Substitution where
-open import Data.Nat
+open import Data.Nat hiding (_⊔_)
open import Data.Fin using (Fin; zero; suc)
open import Data.Vec
open import Function as Fun using (flip)
-open import Data.Star as Star using (Star; ε; _◅_)
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
+ as Star using (Star; ε; _◅_)
+open import Level using (Level; _⊔_)
+import Level as L
+open import Relation.Unary using (Pred)
------------------------------------------------------------------------
-- General functionality
@@ -25,17 +29,17 @@ open import Data.Star as Star using (Star; ε; _◅_)
-- A Sub T m n is a substitution which, when applied to something with
-- at most m variables, yields something with at most n variables.
-Sub : (ℕ → Set) → ℕ → ℕ → Set
+Sub : ∀ {ℓ} → Pred ℕ ℓ → ℕ → ℕ → Set ℓ
Sub T m n = Vec (T n) m
-- A /reversed/ sequence of matching substitutions.
-Subs : (ℕ → Set) → ℕ → ℕ → Set
+Subs : ∀ {ℓ} → Pred ℕ ℓ → ℕ → ℕ → Set ℓ
Subs T = flip (Star (flip (Sub T)))
-- Some simple substitutions.
-record Simple (T : ℕ → Set) : Set where
+record Simple {ℓ : Level} (T : Pred ℕ ℓ) : Set ℓ where
infix 10 _↑
infixl 10 _↑⋆_ _↑✶_
@@ -77,7 +81,8 @@ record Simple (T : ℕ → Set) : Set where
-- Application of substitutions.
-record Application (T₁ T₂ : ℕ → Set) : Set where
+record Application {ℓ₁ ℓ₂ : Level} (T₁ : Pred ℕ ℓ₁) (T₂ : Pred ℕ ℓ₂) :
+ Set (ℓ₁ ⊔ ℓ₂) where
infixl 8 _/_ _/✶_
infixl 9 _⊙_
@@ -97,7 +102,7 @@ record Application (T₁ T₂ : ℕ → Set) : Set where
-- A combination of the two records above.
-record Subst (T : ℕ → Set) : Set where
+record Subst {ℓ : Level} (T : Pred ℕ ℓ) : Set ℓ where
field
simple : Simple T
application : Application T T
@@ -117,7 +122,8 @@ record Subst (T : ℕ → Set) : Set where
-- Liftings from T₁ to T₂.
-record Lift (T₁ T₂ : ℕ → Set) : Set where
+record Lift {ℓ₁ ℓ₂ : Level} (T₁ : Pred ℕ ℓ₁) (T₂ : Pred ℕ ℓ₂) :
+ Set (ℓ₁ ⊔ ℓ₂) where
field
simple : Simple T₁
lift : ∀ {n} → T₁ n → T₂ n
@@ -138,12 +144,12 @@ module VarSubst where
-- "Term" substitutions.
-record TermSubst (T : ℕ → Set) : Set₁ where
+record TermSubst (T : Pred ℕ L.zero) : Set₁ where
field
var : ∀ {n} → Fin n → T n
- app : ∀ {T′} → Lift T′ T → ∀ {m n} → T m → Sub T′ m n → T n
+ app : ∀ {T′ : Pred ℕ L.zero} → Lift T′ T → ∀ {m n} → T m → Sub T′ m n → T n
- module Lifted {T′} (lift : Lift T′ T) where
+ module Lifted {T′ : Pred ℕ L.zero} (lift : Lift T′ T) where
application : Application T T′
application = record { _/_ = app lift }
diff --git a/src/Data/Fin/Substitution/Example.agda b/src/Data/Fin/Substitution/Example.agda
index 737b9ad..9176ebd 100644
--- a/src/Data/Fin/Substitution/Example.agda
+++ b/src/Data/Fin/Substitution/Example.agda
@@ -15,7 +15,8 @@ open import Data.Vec
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl; sym; cong; cong₂)
open PropEq.≡-Reasoning
-open import Data.Star using (Star; ε; _◅_)
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
+ using (Star; ε; _◅_)
-- A representation of the untyped λ-calculus. Uses de Bruijn indices.
@@ -28,7 +29,7 @@ data Tm (n : ℕ) : Set where
-- Code for applying substitutions.
-module TmApp {T} (l : Lift T Tm) where
+module TmApp {ℓ} {T : ℕ → Set ℓ} (l : Lift T Tm) where
open Lift l hiding (var)
-- Applies a substitution to a term.
diff --git a/src/Data/Fin/Substitution/Lemmas.agda b/src/Data/Fin/Substitution/Lemmas.agda
index cf0bb81..9ad89fa 100644
--- a/src/Data/Fin/Substitution/Lemmas.agda
+++ b/src/Data/Fin/Substitution/Lemmas.agda
@@ -8,15 +8,18 @@ module Data.Fin.Substitution.Lemmas where
import Category.Applicative.Indexed as Applicative
open import Data.Fin.Substitution
-open import Data.Nat
+open import Data.Nat hiding (_⊔_)
open import Data.Fin using (Fin; zero; suc; lift)
open import Data.Vec
import Data.Vec.Properties as VecProp
open import Function as Fun using (_∘_; _$_)
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl; sym; cong; cong₂)
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
+ using (Star; ε; _◅_; _▻_)
open PropEq.≡-Reasoning
-open import Data.Star using (Star; ε; _◅_; _▻_)
+open import Level using (Level; _⊔_)
+open import Relation.Unary using (Pred)
-- A lemma which does not refer to any substitutions.
@@ -31,7 +34,7 @@ lift-commutes k (suc j) (suc x) = cong suc (lift-commutes k j x)
-- assumption that the underlying substitution machinery satisfies
-- certain properties.
-record Lemmas₀ (T : ℕ → Set) : Set where
+record Lemmas₀ {ℓ : Level} (T : Pred ℕ ℓ) : Set ℓ where
field simple : Simple T
open Simple simple
@@ -60,7 +63,7 @@ record Lemmas₀ (T : ℕ → Set) : Set where
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
+record Lemmas₁ {ℓ} (T : Pred ℕ ℓ) : Set ℓ where
field lemmas₀ : Lemmas₀ T
open Lemmas₀ lemmas₀
@@ -117,7 +120,7 @@ record Lemmas₁ (T : ℕ → Set) : Set where
open Lemmas₀ lemmas₀ public
-record Lemmas₂ (T : ℕ → Set) : Set where
+record Lemmas₂ {ℓ} (T : Pred ℕ ℓ) : Set ℓ where
field
lemmas₁ : Lemmas₁ T
application : Application T T
@@ -202,7 +205,7 @@ record Lemmas₂ (T : ℕ → Set) : Set where
open Subst subst public hiding (simple; application)
open Lemmas₁ lemmas₁ public
-record Lemmas₃ (T : ℕ → Set) : Set where
+record Lemmas₃ {ℓ} (T : Pred ℕ ℓ) : Set ℓ where
field lemmas₂ : Lemmas₂ T
open Lemmas₂ lemmas₂
@@ -232,7 +235,7 @@ record Lemmas₃ (T : ℕ → Set) : Set where
open Lemmas₂ lemmas₂ public hiding (wk-⊙-sub′)
-record Lemmas₄ (T : ℕ → Set) : Set where
+record Lemmas₄ {ℓ} (T : Pred ℕ ℓ) : Set ℓ where
field lemmas₃ : Lemmas₃ T
open Lemmas₃ lemmas₃
@@ -348,7 +351,7 @@ record Lemmas₄ (T : ℕ → Set) : Set where
-- For an example of how AppLemmas can be used, see
-- Data.Fin.Substitution.List.
-record AppLemmas (T₁ T₂ : ℕ → Set) : Set where
+record AppLemmas {ℓ₁ ℓ₂} (T₁ : Pred ℕ ℓ₁) (T₂ : Pred ℕ ℓ₂) : Set (ℓ₁ ⊔ ℓ₂) where
field
application : Application T₁ T₂
lemmas₄ : Lemmas₄ T₂
@@ -401,7 +404,7 @@ record AppLemmas (T₁ T₂ : ℕ → Set) : Set where
hiding (application; _⊙_; _/_; _/✶_;
id-vanishes; /-⊙; wk-commutes)
-record Lemmas₅ (T : ℕ → Set) : Set where
+record Lemmas₅ {ℓ} (T : Pred ℕ ℓ) : Set ℓ where
field lemmas₄ : Lemmas₄ T
private module L₄ = Lemmas₄ lemmas₄
diff --git a/src/Data/Fin/Substitution/List.agda b/src/Data/Fin/Substitution/List.agda
index ba5b548..b24ccbd 100644
--- a/src/Data/Fin/Substitution/List.agda
+++ b/src/Data/Fin/Substitution/List.agda
@@ -8,8 +8,9 @@
-- can be used.
open import Data.Fin.Substitution.Lemmas
+open import Data.Nat using (ℕ)
-module Data.Fin.Substitution.List {T} (lemmas₄ : Lemmas₄ T) where
+module Data.Fin.Substitution.List {ℓ} {T : ℕ → Set ℓ} (lemmas₄ : Lemmas₄ T) where
open import Data.List.Base
open import Data.List.Properties
diff --git a/src/Data/Float.agda b/src/Data/Float.agda
index 77af9ce..1c11919 100644
--- a/src/Data/Float.agda
+++ b/src/Data/Float.agda
@@ -6,23 +6,20 @@
module Data.Float where
-open import Data.Bool.Base using (Bool; false; true)
-open import Relation.Nullary.Decidable
-open import Relation.Nullary
-open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
-open import Relation.Binary.PropositionalEquality.TrustMe
open import Data.String.Base using (String)
+------------------------------------------------------------------------
+-- Re-export built-ins publically
+
open import Agda.Builtin.Float public
- using ( Float; primFloatEquality; primShowFloat )
+ using
+ ( Float
+ ; primFloatEquality
+ ; primShowFloat
+ )
+
+------------------------------------------------------------------------
+-- Operations
show : Float → String
show = primShowFloat
-
-infix 4 _≟_
-
-_≟_ : (x y : Float) → Dec (x ≡ y)
-x ≟ y with primFloatEquality x y
-... | true = yes trustMe
-... | false = no whatever
- where postulate whatever : _
diff --git a/src/Data/Float/Unsafe.agda b/src/Data/Float/Unsafe.agda
new file mode 100644
index 0000000..de205d2
--- /dev/null
+++ b/src/Data/Float/Unsafe.agda
@@ -0,0 +1,24 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Unsafe Float operations
+------------------------------------------------------------------------
+
+module Data.Float.Unsafe where
+
+open import Data.Float
+open import Data.Bool.Base using (false; true)
+open import Relation.Nullary using (Dec; yes; no)
+open import Relation.Binary.PropositionalEquality using (_≡_)
+open import Relation.Binary.PropositionalEquality.TrustMe
+
+------------------------------------------------------------------------
+-- Equality testing
+
+infix 4 _≟_
+
+_≟_ : (x y : Float) → Dec (x ≡ y)
+x ≟ y with primFloatEquality x y
+... | true = yes trustMe
+... | false = no whatever
+ where postulate whatever : _
diff --git a/src/Data/Graph/Acyclic.agda b/src/Data/Graph/Acyclic.agda
index 3a93790..4aaa7e6 100644
--- a/src/Data/Graph/Acyclic.agda
+++ b/src/Data/Graph/Acyclic.agda
@@ -10,11 +10,13 @@
module Data.Graph.Acyclic where
+open import Level using (_⊔_)
open import Data.Nat.Base as Nat using (ℕ; zero; suc; _<′_)
import Data.Nat.Properties as Nat
open import Data.Fin as Fin
- using (Fin; Fin′; zero; suc; #_; toℕ) renaming (_ℕ-ℕ_ to _-_)
-open import Data.Fin.Properties as FP using (_≟_)
+ using (Fin; Fin′; zero; suc; #_; toℕ; _≟_) renaming (_ℕ-ℕ_ to _-_)
+import Data.Fin.Properties as FP
+import Data.Fin.Permutation.Components as PC
open import Data.Product as Prod using (∃; _×_; _,_)
open import Data.Maybe.Base using (Maybe; nothing; just)
open import Data.Empty
@@ -38,7 +40,7 @@ private
------------------------------------------------------------------------
-- Node contexts
-record Context (Node Edge : Set) (n : ℕ) : Set where
+record Context {ℓ e} (Node : Set ℓ) (Edge : Set e) (n : ℕ) : Set (ℓ ⊔ e) where
constructor context
field
label : Node
@@ -48,10 +50,12 @@ open Context public
-- Map for contexts.
-cmap : ∀ {N₁ N₂ E₁ E₂ n} →
- (N₁ → N₂) → (List (E₁ × Fin n) → List (E₂ × Fin n)) →
- Context N₁ E₁ n → Context N₂ E₂ n
-cmap f g c = context (f (label c)) (g (successors c))
+module _ {ℓ₁ e₁} {N₁ : Set ℓ₁} {E₁ : Set e₁}
+ {ℓ₂ e₂} {N₂ : Set ℓ₂} {E₂ : Set e₂} where
+
+ cmap : ∀ {n} → (N₁ → N₂) → (List (E₁ × Fin n) → List (E₂ × Fin n)) →
+ Context N₁ E₁ n → Context N₂ E₂ n
+ cmap f g c = context (f (label c)) (g (successors c))
------------------------------------------------------------------------
-- Graphs
@@ -60,7 +64,7 @@ infixr 3 _&_
-- The DAGs are indexed on the number of nodes.
-data Graph (Node Edge : Set) : ℕ → Set where
+data Graph {ℓ e} (Node : Set ℓ) (Edge : Set e) : ℕ → Set (ℓ ⊔ e) where
∅ : Graph Node Edge 0
_&_ : ∀ {n} (c : Context Node Edge n) (g : Graph Node Edge n) →
Graph Node Edge (suc n)
@@ -78,48 +82,54 @@ private
------------------------------------------------------------------------
-- Higher-order functions
+module _ {ℓ e} {N : Set ℓ} {E : Set e} {t} where
+
-- "Fold right".
-foldr : ∀ {N E m} (T : ℕ → Set) →
- (∀ {n} → Context N E n → T n → T (suc n)) →
- T 0 →
- Graph N E m → T m
-foldr T _∙_ x ∅ = x
-foldr T _∙_ x (c & g) = c ∙ foldr T _∙_ x g
+ foldr : (T : ℕ → Set t) →
+ (∀ {n} → Context N E n → T n → T (suc n)) →
+ T 0 →
+ ∀ {m} → Graph N E m → T m
+ foldr T _∙_ x ∅ = x
+ foldr T _∙_ x (c & g) = c ∙ foldr T _∙_ x g
-- "Fold left".
-foldl : ∀ {N E n} (T : ℕ → Set) →
- ((i : Fin n) → T (toℕ i) → Context N E (n - suc i) →
- T (suc (toℕ i))) →
- T 0 →
- Graph N E n → T n
-foldl T f x ∅ = x
-foldl T f x (c & g) =
- foldl (λ n → T (suc n)) (λ i → f (suc i)) (f zero x c) g
+ foldl : ∀ {n} (T : ℕ → Set t) →
+ ((i : Fin n) → T (toℕ i) → Context N E (n - suc i) →
+ T (suc (toℕ i))) →
+ T 0 →
+ Graph N E n → T n
+ foldl T f x ∅ = x
+ foldl T f x (c & g) = foldl (T ∘′ suc) (f ∘ suc) (f zero x c) g
+
+
+module _ {ℓ₁ e₁} {N₁ : Set ℓ₁} {E₁ : Set e₁}
+ {ℓ₂ e₂} {N₂ : Set ℓ₂} {E₂ : Set e₂} where
-- Maps over node contexts.
-map : ∀ {N₁ N₂ E₁ E₂ n} →
- (∀ {n} → Context N₁ E₁ n → Context N₂ E₂ n) →
- Graph N₁ E₁ n → Graph N₂ E₂ n
-map f = foldr _ (λ c g → f c & g) ∅
+ map : (∀ {n} → Context N₁ E₁ n → Context N₂ E₂ n) →
+ ∀ {n} → Graph N₁ E₁ n → Graph N₂ E₂ n
+ map f = foldr _ (λ c → f c &_) ∅
-- Maps over node labels.
-nmap : ∀ {N₁ N₂ E n} → (N₁ → N₂) → Graph N₁ E n → Graph N₂ E n
+nmap : ∀ {ℓ₁ ℓ₂ e} {N₁ : Set ℓ₁} {N₂ : Set ℓ₂} {E : Set e} →
+ ∀ {n} → (N₁ → N₂) → Graph N₁ E n → Graph N₂ E n
nmap f = map (cmap f id)
-- Maps over edge labels.
-emap : ∀ {N E₁ E₂ n} → (E₁ → E₂) → Graph N E₁ n → Graph N E₂ n
+emap : ∀ {ℓ e₁ e₂} {N : Set ℓ} {E₁ : Set e₁} {E₂ : Set e₂} →
+ ∀ {n} → (E₁ → E₂) → Graph N E₁ n → Graph N E₂ n
emap f = map (cmap id (List.map (Prod.map f id)))
-- Zips two graphs with the same number of nodes. Note that one of the
-- graphs has a type which restricts it to be completely disconnected.
-zipWith : ∀ {N₁ N₂ N E n} → (N₁ → N₂ → N) →
- Graph N₁ ⊥ n → Graph N₂ E n → Graph N E n
+zipWith : ∀ {ℓ₁ ℓ₂ ℓ e} {N₁ : Set ℓ₁} {N₂ : Set ℓ₂} {N : Set ℓ} {E : Set e} →
+ ∀ {n} → (N₁ → N₂ → N) → Graph N₁ ⊥ n → Graph N₂ E n → Graph N E n
zipWith _∙_ ∅ ∅ = ∅
zipWith _∙_ (c₁ & g₁) (c₂ & g₂) =
context (label c₁ ∙ label c₂) (successors c₂) & zipWith _∙_ g₁ g₂
@@ -138,37 +148,38 @@ disconnected (suc n) = context tt [] & disconnected n
complete : ∀ n → Graph ⊤ ⊤ n
complete zero = ∅
complete (suc n) =
- context tt (List.map (_,_ tt) $ Vec.toList (Vec.allFin n)) &
+ context tt (List.map (tt ,_) $ Vec.toList (Vec.allFin n)) &
complete n
------------------------------------------------------------------------
-- Queries
+module _ {ℓ e} {N : Set ℓ} {E : Set e} where
+
-- The top-most context.
-head : ∀ {N E n} → Graph N E (suc n) → Context N E n
-head (c & g) = c
+ head : ∀ {n} → Graph N E (suc n) → Context N E n
+ head (c & g) = c
-- The remaining graph.
-tail : ∀ {N E n} → Graph N E (suc n) → Graph N E n
-tail (c & g) = g
+ tail : ∀ {n} → Graph N E (suc n) → Graph N E n
+ tail (c & g) = g
-- Finds the context and remaining graph corresponding to a given node
-- index.
-_[_] : ∀ {N E n} →
- Graph N E n → (i : Fin n) → Graph N E (suc (n - suc i))
-∅ [ () ]
-(c & g) [ zero ] = c & g
-(c & g) [ suc i ] = g [ i ]
+ _[_] : ∀ {n} → Graph N E n → (i : Fin n) → Graph N E (suc (n - suc i))
+ ∅ [ () ]
+ (c & g) [ zero ] = c & g
+ (c & g) [ suc i ] = g [ i ]
-- The nodes of the graph (node number relative to "topmost" node ×
-- node label).
-nodes : ∀ {N E n} → Graph N E n → Vec (Fin n × N) n
-nodes {N} = Vec.zip (Vec.allFin _) ∘
- foldr (Vec N) (λ c ns → label c ∷ ns) []
+ nodes : ∀ {n} → Graph N E n → Vec (Fin n × N) n
+ nodes = Vec.zip (Vec.allFin _) ∘
+ foldr (Vec N) (λ c → label c ∷_) []
private
@@ -176,11 +187,14 @@ private
(# 3 , 3) ∷ (# 4 , 4) ∷ []
test-nodes = P.refl
+
+module _ {ℓ e} {N : Set ℓ} {E : Set e} where
+
-- Topological sort. Gives a vector where earlier nodes are never
-- successors of later nodes.
-topSort : ∀ {N E n} → Graph N E n → Vec (Fin n × N) n
-topSort = nodes
+ topSort : ∀ {n} → Graph N E n → Vec (Fin n × N) n
+ topSort = nodes
-- The edges of the graph (predecessor × edge label × successor).
--
@@ -188,11 +202,11 @@ topSort = nodes
-- the graph, and the successor is a node number relative to the
-- predecessor.
-edges : ∀ {E N n} → Graph N E n → List (∃ λ i → E × Fin (n - suc i))
-edges {E} {N} {n} =
- foldl (λ _ → List (∃ λ i → E × Fin (n - suc i)))
- (λ i es c → List._++_ es (List.map (_,_ i) (successors c)))
- []
+ edges : ∀ {n} → Graph N E n → List (∃ λ i → E × Fin (n - suc i))
+ edges {n} =
+ foldl (λ _ → List (∃ λ i → E × Fin (n - suc i)))
+ (λ i es c → es List.++ List.map (i ,_) (successors c))
+ []
private
@@ -203,8 +217,8 @@ private
-- The successors of a given node i (edge label × node number relative
-- to i).
-sucs : ∀ {E N n} →
- Graph N E n → (i : Fin n) → List (E × Fin (n - suc i))
+sucs : ∀ {ℓ e} {N : Set ℓ} {E : Set e} →
+ ∀ {n} → Graph N E n → (i : Fin n) → List (E × Fin (n - suc i))
sucs g i = successors $ head (g [ i ])
private
@@ -215,13 +229,14 @@ private
-- The predecessors of a given node i (node number relative to i ×
-- edge label).
-preds : ∀ {E N n} → Graph N E n → (i : Fin n) → List (Fin′ i × E)
+preds : ∀ {ℓ e} {N : Set ℓ} {E : Set e} →
+ ∀ {n} → Graph N E n → (i : Fin n) → List (Fin′ i × E)
preds g zero = []
preds (c & g) (suc i) =
- List._++_ (List.gfilter (p i) $ successors c)
+ List._++_ (List.mapMaybe (p i) $ successors c)
(List.map (Prod.map suc id) $ preds g i)
where
- p : ∀ {E : Set} {n} (i : Fin n) → E × Fin n → Maybe (Fin′ (suc i) × E)
+ p : ∀ {e} {E : Set e} {n} (i : Fin n) → E × Fin n → Maybe (Fin′ (suc i) × E)
p i (e , j) with i ≟ j
p i (e , .i) | yes P.refl = just (zero , e)
p i (e , j) | no _ = nothing
@@ -242,10 +257,11 @@ weaken {n} {i} j = Fin.inject≤ j (FP.nℕ-ℕi≤n n (suc i))
-- Labels each node with its node number.
-number : ∀ {N E n} → Graph N E n → Graph (Fin n × N) E n
-number {N} {E} =
+number : ∀ {ℓ e} {N : Set ℓ} {E : Set e} →
+ ∀ {n} → Graph N E n → Graph (Fin n × N) E n
+number {N = N} {E} =
foldr (λ n → Graph (Fin n × N) E n)
- (λ c g → cmap (_,_ zero) id c & nmap (Prod.map suc id) g)
+ (λ c g → cmap (zero ,_) id c & nmap (Prod.map suc id) g)
private
@@ -261,12 +277,13 @@ private
-- Reverses all the edges in the graph.
-reverse : ∀ {N E n} → Graph N E n → Graph N E n
-reverse {N} {E} g =
+reverse : ∀ {ℓ e} {N : Set ℓ} {E : Set e} →
+ ∀ {n} → Graph N E n → Graph N E n
+reverse {N = N} {E} g =
foldl (Graph N E)
(λ i g' c →
context (label c)
- (List.map (Prod.swap ∘ Prod.map FP.reverse id) $
+ (List.map (Prod.swap ∘ Prod.map PC.reverse id) $
preds g i)
& g')
∅ g
@@ -282,25 +299,27 @@ private
-- Expands the subgraph induced by a given node into a tree (thus
-- losing all sharing).
-data Tree (N E : Set) : Set where
+data Tree {ℓ e} (N : Set ℓ) (E : Set e) : Set (ℓ ⊔ e) 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 ])
- where
- Pred = λ n → Graph N E (suc n) → Tree N E
+module _ {ℓ e} {N : Set ℓ} {E : Set e} where
+
+ toTree : ∀ {n} → Graph N E n → Fin n → Tree N E
+ toTree 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 (c & g) =
- node (label c)
- (List.map
- (Prod.map id (λ i → rec (n - suc i) (lemma n i) (g [ i ])))
- (successors c))
+ expand : (n : ℕ) → <′-Rec Pred n → Pred n
+ expand n rec (c & g) =
+ node (label c)
+ (List.map
+ (Prod.map id (λ i → rec (n - suc i) (lemma n i) (g [ i ])))
+ (successors c))
-- Performs the toTree expansion once for each node.
-toForest : ∀ {N E n} → Graph N E n → Vec (Tree N E) n
-toForest g = Vec.map (toTree g) (Vec.allFin _)
+ toForest : ∀ {n} → Graph N E n → Vec (Tree N E) n
+ toForest g = Vec.map (toTree g) (Vec.allFin _)
private
diff --git a/src/Data/Integer.agda b/src/Data/Integer.agda
index cc6b827..d281d08 100644
--- a/src/Data/Integer.agda
+++ b/src/Data/Integer.agda
@@ -6,19 +6,9 @@
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; _++_)
-open import Function
-open import Data.Sum
-open import Relation.Nullary
-import Relation.Nullary.Decidable as Dec
-open import Relation.Binary
-open import Relation.Binary.PropositionalEquality
- using (_≡_; refl; sym; subst; cong; cong₂; module ≡-Reasoning)
-open ≡-Reasoning
------------------------------------------------------------------------
-- Integers, basic types and operations
@@ -26,9 +16,13 @@ open ≡-Reasoning
open import Data.Integer.Base public
------------------------------------------------------------------------
--- Conversions
+-- Re-export queries from the properties modules
+
+open import Data.Integer.Properties public
+ using (_≟_; _≤?_)
--- Decimal string representation.
+------------------------------------------------------------------------
+-- Conversions
show : ℤ → String
show i = showSign (sign i) ++ ℕ.show ∣ i ∣
@@ -38,36 +32,10 @@ show i = showSign (sign i) ++ ℕ.show ∣ i ∣
showSign Sign.+ = ""
------------------------------------------------------------------------
--- Properties of the view of integers as sign + absolute value
-
-◃-cong : ∀ {i j} → sign i ≡ sign j → ∣ i ∣ ≡ ∣ j ∣ → i ≡ j
-◃-cong {i} {j} sign-≡ abs-≡ = begin
- i ≡⟨ sym $ ◃-left-inverse i ⟩
- sign i ◃ ∣ i ∣ ≡⟨ cong₂ _◃_ sign-≡ abs-≡ ⟩
- sign j ◃ ∣ j ∣ ≡⟨ ◃-left-inverse j ⟩
- j ∎
-
-signAbs : ∀ i → SignAbs i
-signAbs i = subst SignAbs (◃-left-inverse i) (sign i ◂ ∣ i ∣)
-
-------------------------------------------------------------------------
--- Equality is decidable
-
-infix 4 _≟_
-
-_≟_ : Decidable {A = ℤ} _≡_
-i ≟ j with Sign._≟_ (sign i) (sign j) | ℕ._≟_ ∣ i ∣ ∣ j ∣
-i ≟ j | yes sign-≡ | yes abs-≡ = yes (◃-cong sign-≡ abs-≡)
-i ≟ j | no sign-≢ | _ = no (sign-≢ ∘ cong sign)
-i ≟ j | _ | no abs-≢ = no (abs-≢ ∘ cong ∣_∣)
-
-------------------------------------------------------------------------
--- Ordering is decidable
+-- Deprecated
-infix 4 _≤?_
+-- Version 0.17
-_≤?_ : Decidable _≤_
--[1+ m ] ≤? -[1+ n ] = Dec.map′ -≤- drop‿-≤- (ℕ._≤?_ n m)
--[1+ m ] ≤? + n = yes -≤+
-+ m ≤? -[1+ n ] = no λ ()
-+ m ≤? + n = Dec.map′ +≤+ drop‿+≤+ (ℕ._≤?_ m n)
+open import Data.Integer.Properties public
+ using (◃-cong; drop‿+≤+; drop‿-≤-)
+ renaming (◃-inverse to ◃-left-inverse)
diff --git a/src/Data/Integer/Addition/Properties.agda b/src/Data/Integer/Addition/Properties.agda
index 2da1bdf..4dc347e 100644
--- a/src/Data/Integer/Addition/Properties.agda
+++ b/src/Data/Integer/Addition/Properties.agda
@@ -18,8 +18,8 @@ open import Data.Integer.Properties public
)
renaming
( +-comm to comm
- ; +-identityˡ to identityˡ
- ; +-identityʳ to identityʳ
+ ; +-identityˡ to identityˡ
+ ; +-identityʳ to identityʳ
; +-assoc to assoc
; +-isSemigroup to isSemigroup
; +-0-isCommutativeMonoid to isCommutativeMonoid
diff --git a/src/Data/Integer/Base.agda b/src/Data/Integer/Base.agda
index 5439c63..c943b4c 100644
--- a/src/Data/Integer/Base.agda
+++ b/src/Data/Integer/Base.agda
@@ -10,10 +10,9 @@ open import Data.Nat.Base as ℕ
using (ℕ) renaming (_+_ to _ℕ+_; _*_ to _ℕ*_)
open import Data.Sign as Sign using (Sign) renaming (_*_ to _S*_)
open import Function
-open import Relation.Nullary
-open import Relation.Binary
-open import Relation.Binary.Core using (_≡_; refl)
--- Importing Core here ^^^ to keep a small import list
+open import Relation.Nullary using (¬_)
+open import Relation.Binary using (Rel)
+open import Relation.Binary.PropositionalEquality using (_≡_)
infix 8 -_
infixl 7 _*_ _⊓_
@@ -23,13 +22,13 @@ infix 4 _≤_ _≥_ _<_ _>_ _≰_ _≱_ _≮_ _≯_
------------------------------------------------------------------------
-- The types
--- Integers.
-
open import Agda.Builtin.Int public
using ()
- renaming ( Int to ℤ
- ; negsuc to -[1+_] -- -[1+ n ] stands for - (1 + n).
- ; pos to +_ ) -- + n stands for n.
+ renaming
+ ( Int to ℤ
+ ; pos to +_ -- "+ n" stands for "n"
+ ; negsuc to -[1+_] -- "-[1+ n ]" stands for "- (1 + n)"
+ )
------------------------------------------------------------------------
-- Conversions
@@ -56,14 +55,14 @@ _ ◃ ℕ.zero = + ℕ.zero
Sign.+ ◃ n = + n
Sign.- ◃ ℕ.suc n = -[1+ n ]
-◃-left-inverse : ∀ i → sign i ◃ ∣ i ∣ ≡ i
-◃-left-inverse -[1+ n ] = refl
-◃-left-inverse (+ ℕ.zero) = refl
-◃-left-inverse (+ ℕ.suc n) = refl
-
data SignAbs : ℤ → Set where
_◂_ : (s : Sign) (n : ℕ) → SignAbs (s ◃ n)
+signAbs : ∀ i → SignAbs i
+signAbs (+ ℕ.zero) = Sign.+ ◂ ℕ.zero
+signAbs (+ (ℕ.suc n)) = Sign.+ ◂ ℕ.suc n
+signAbs (-[1+ n ]) = Sign.- ◂ ℕ.suc n
+
------------------------------------------------------------------------
-- Arithmetic
@@ -92,33 +91,17 @@ _+_ : ℤ → ℤ → ℤ
-- Subtraction.
_-_ : ℤ → ℤ → ℤ
-i - j = i + - j
+i - j = i + (- j)
-- Successor.
suc : ℤ → ℤ
-suc i = + 1 + i
-
-private
-
- suc-is-lazy⁺ : ∀ n → suc (+ n) ≡ + ℕ.suc n
- suc-is-lazy⁺ n = refl
-
- suc-is-lazy⁻ : ∀ n → suc -[1+ ℕ.suc n ] ≡ -[1+ n ]
- suc-is-lazy⁻ n = refl
+suc i = (+ 1) + i
-- Predecessor.
pred : ℤ → ℤ
-pred i = - + 1 + i
-
-private
-
- pred-is-lazy⁺ : ∀ n → pred (+ ℕ.suc n) ≡ + n
- pred-is-lazy⁺ n = refl
-
- pred-is-lazy⁻ : ∀ n → pred -[1+ n ] ≡ -[1+ ℕ.suc n ]
- pred-is-lazy⁻ n = refl
+pred i = (- + 1) + i
-- Multiplication.
@@ -169,9 +152,3 @@ x ≮ y = ¬ (x < y)
_≯_ : Rel ℤ _
x ≯ y = ¬ (x > y)
-
-drop‿+≤+ : ∀ {m n} → + m ≤ + n → ℕ._≤_ m n
-drop‿+≤+ (+≤+ m≤n) = m≤n
-
-drop‿-≤- : ∀ {m n} → -[1+ m ] ≤ -[1+ n ] → ℕ._≤_ n m
-drop‿-≤- (-≤- n≤m) = n≤m
diff --git a/src/Data/Integer/Divisibility.agda b/src/Data/Integer/Divisibility.agda
index bb3b23c..63b4fca 100644
--- a/src/Data/Integer/Divisibility.agda
+++ b/src/Data/Integer/Divisibility.agda
@@ -31,4 +31,4 @@ Coprime = ℕ.Coprime on ∣_∣
coprime-divisor : ∀ i j k → Coprime i j → i ∣ j * k → i ∣ k
coprime-divisor i j k c eq =
- ℕ.coprime-divisor c (subst (ℕ._∣_ ∣ i ∣) (abs-*-commute j k) eq)
+ ℕ.coprime-divisor c (subst (∣ i ∣ ℕ.∣_ ) (abs-*-commute j k) eq)
diff --git a/src/Data/Integer/Literals.agda b/src/Data/Integer/Literals.agda
new file mode 100644
index 0000000..3c6149d
--- /dev/null
+++ b/src/Data/Integer/Literals.agda
@@ -0,0 +1,24 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Integer Literals
+------------------------------------------------------------------------
+
+module Data.Integer.Literals where
+
+open import Agda.Builtin.FromNat
+open import Agda.Builtin.FromNeg
+open import Data.Unit
+open import Data.Integer
+
+number : Number ℤ
+number = record
+ { Constraint = λ _ → ⊤
+ ; fromNat = λ n → + n
+ }
+
+negative : Negative ℤ
+negative = record
+ { Constraint = λ _ → ⊤
+ ; fromNeg = λ n → - (+ n)
+ }
diff --git a/src/Data/Integer/Properties.agda b/src/Data/Integer/Properties.agda
index 543e2b7..31f4af8 100644
--- a/src/Data/Integer/Properties.agda
+++ b/src/Data/Integer/Properties.agda
@@ -7,18 +7,16 @@
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 renaming (suc to sucℤ)
-open import Data.Nat
- using (ℕ; suc; zero; _∸_; s≤s; z≤n; ≤-pred)
+open import Data.Integer.Base renaming (suc to sucℤ)
+open import Data.Nat as ℕ
+ using (ℕ; suc; zero; _∸_; s≤s; z≤n)
hiding (module ℕ)
renaming (_+_ to _ℕ+_; _*_ to _ℕ*_;
- _<_ to _ℕ<_; _≥_ to _ℕ≥_; _≰_ to _ℕ≰_; _≤?_ to _ℕ≤?_)
+ _≤_ to _ℕ≤_; _<_ to _ℕ<_; _≥_ to _ℕ≥_; _≰_ to _ℕ≰_; _≟_ to _ℕ≟_; _≤?_ to _ℕ≤?_)
import Data.Nat.Properties as ℕₚ
+open import Data.Nat.Solver
open import Data.Product using (proj₁; proj₂; _,_)
open import Data.Sum using (inj₁; inj₂)
open import Data.Sign as Sign using () renaming (_*_ to _𝕊*_)
@@ -26,14 +24,17 @@ import Data.Sign.Properties as 𝕊ₚ
open import Function using (_∘_; _$_)
open import Relation.Binary
open import Relation.Binary.PropositionalEquality
+import Relation.Binary.PartialOrderReasoning as POR
open import Relation.Nullary using (yes; no)
open import Relation.Nullary.Negation using (contradiction)
+import Relation.Nullary.Decidable as Dec
-open Algebra.FunctionProperties (_≡_ {A = ℤ})
-open Algebra.FunctionProperties.Consequences (setoid ℤ)
+open import Algebra.FunctionProperties (_≡_ {A = ℤ})
+open import Algebra.FunctionProperties.Consequences (setoid ℤ)
+open import Algebra.Structures (_≡_ {A = ℤ})
open Morphism.Definitions ℤ ℕ _≡_
-open ℕₚ.SemiringSolver
open ≡-Reasoning
+open +-*-Solver
------------------------------------------------------------------------
-- Equality
@@ -44,22 +45,29 @@ open ≡-Reasoning
-[1+-injective : ∀ {m n} → -[1+ m ] ≡ -[1+ n ] → m ≡ n
-[1+-injective refl = refl
+infix 4 _≟_
+_≟_ : Decidable {A = ℤ} _≡_
++ m ≟ + n = Dec.map′ (cong (+_)) +-injective (m ℕ≟ n)
++ m ≟ -[1+ n ] = no λ()
+-[1+ m ] ≟ + n = no λ()
+-[1+ m ] ≟ -[1+ n ] = Dec.map′ (cong -[1+_]) -[1+-injective (m ℕ≟ n)
+
≡-decSetoid : DecSetoid _ _
≡-decSetoid = decSetoid _≟_
------------------------------------------------------------------------
-- Properties of -_
-doubleNeg : ∀ n → - - n ≡ n
-doubleNeg (+ zero) = refl
-doubleNeg (+ suc n) = refl
-doubleNeg (-[1+ n ]) = refl
+neg-involutive : ∀ n → - - n ≡ n
+neg-involutive (+ zero) = refl
+neg-involutive (+ suc n) = refl
+neg-involutive (-[1+ n ]) = refl
neg-injective : ∀ {m n} → - m ≡ - n → m ≡ n
neg-injective {m} {n} -m≡-n = begin
- m ≡⟨ sym (doubleNeg m) ⟩
+ m ≡⟨ sym (neg-involutive m) ⟩
- - m ≡⟨ cong -_ -m≡-n ⟩
- - - n ≡⟨ doubleNeg n ⟩
+ - - n ≡⟨ neg-involutive n ⟩
n ∎
------------------------------------------------------------------------
@@ -77,6 +85,18 @@ neg-injective {m} {n} -m≡-n = begin
------------------------------------------------------------------------
-- Properties of sign and _◃_
+◃-inverse : ∀ i → sign i ◃ ∣ i ∣ ≡ i
+◃-inverse -[1+ n ] = refl
+◃-inverse (+ zero) = refl
+◃-inverse (+ suc n) = refl
+
+◃-cong : ∀ {i j} → sign i ≡ sign j → ∣ i ∣ ≡ ∣ j ∣ → i ≡ j
+◃-cong {i} {j} sign-≡ abs-≡ = begin
+ i ≡⟨ sym $ ◃-inverse i ⟩
+ sign i ◃ ∣ i ∣ ≡⟨ cong₂ _◃_ sign-≡ abs-≡ ⟩
+ sign j ◃ ∣ j ∣ ≡⟨ ◃-inverse j ⟩
+ j ∎
+
+◃n≡+n : ∀ n → Sign.+ ◃ n ≡ + n
+◃n≡+n zero = refl
+◃n≡+n (suc _) = refl
@@ -106,8 +126,7 @@ sign-cong {s₁} {s₂} {n₁} {n₂} eq = begin
sign (s₂ ◃ suc n₂) ≡⟨ sign-◃ s₂ n₂ ⟩
s₂ ∎
-abs-cong : ∀ {s₁ s₂ n₁ n₂} →
- s₁ ◃ n₁ ≡ s₂ ◃ n₂ → n₁ ≡ n₂
+abs-cong : ∀ {s₁ s₂ n₁ n₂} → s₁ ◃ n₁ ≡ s₂ ◃ n₂ → n₁ ≡ n₂
abs-cong {s₁} {s₂} {n₁} {n₂} eq = begin
n₁ ≡⟨ sym $ abs-◃ s₁ n₁ ⟩
∣ s₁ ◃ n₁ ∣ ≡⟨ cong ∣_∣ eq ⟩
@@ -161,16 +180,52 @@ sign-⊖-≰ = sign-⊖-< ∘ ℕₚ.≰⇒>
-[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
+∣m⊖n∣≡∣n⊖m∣ : ∀ x y → ∣ x ⊖ y ∣ ≡ ∣ y ⊖ x ∣
+∣m⊖n∣≡∣n⊖m∣ zero zero = refl
+∣m⊖n∣≡∣n⊖m∣ zero (suc _) = refl
+∣m⊖n∣≡∣n⊖m∣ (suc _) zero = refl
+∣m⊖n∣≡∣n⊖m∣ (suc x) (suc y) = ∣m⊖n∣≡∣n⊖m∣ x y
+
++-cancelˡ-⊖ : ∀ a b c → (a ℕ+ b) ⊖ (a ℕ+ c) ≡ b ⊖ c
++-cancelˡ-⊖ zero b c = refl
++-cancelˡ-⊖ (suc a) b c = +-cancelˡ-⊖ a b c
+
+------------------------------------------------------------------------
+-- Properties of _-_
+
+neg-minus-pos : ∀ x y → -[1+ x ] - (+ y) ≡ -[1+ (y ℕ+ x) ]
+neg-minus-pos x zero = refl
+neg-minus-pos zero (suc y) = cong (-[1+_] ∘ suc) (sym (ℕₚ.+-identityʳ y))
+neg-minus-pos (suc x) (suc y) = cong (-[1+_] ∘ suc) (ℕₚ.+-comm (suc x) y)
+
+[+m]-[+n]≡m⊖n : ∀ x y → (+ x) - (+ y) ≡ x ⊖ y
+[+m]-[+n]≡m⊖n zero zero = refl
+[+m]-[+n]≡m⊖n zero (suc y) = refl
+[+m]-[+n]≡m⊖n (suc x) zero = cong (+_ ∘ suc) (ℕₚ.+-identityʳ x)
+[+m]-[+n]≡m⊖n (suc x) (suc y) = refl
+
+∣m-n∣≡∣n-m∣ : (x y : ℤ) → ∣ x - y ∣ ≡ ∣ y - x ∣
+∣m-n∣≡∣n-m∣ -[1+ x ] -[1+ y ] = ∣m⊖n∣≡∣n⊖m∣ y x
+∣m-n∣≡∣n-m∣ -[1+ x ] (+ y) = begin
+ ∣ -[1+ x ] - (+ y) ∣ ≡⟨ cong ∣_∣ (neg-minus-pos x y) ⟩
+ suc (y ℕ+ x) ≡⟨ sym (ℕₚ.+-suc y x) ⟩
+ y ℕ+ suc x ∎
+∣m-n∣≡∣n-m∣ (+ x) -[1+ y ] = begin
+ x ℕ+ suc y ≡⟨ ℕₚ.+-suc x y ⟩
+ suc (x ℕ+ y) ≡⟨ cong ∣_∣ (sym (neg-minus-pos y x)) ⟩
+ ∣ -[1+ y ] + - (+ x) ∣ ∎
+∣m-n∣≡∣n-m∣ (+ x) (+ y) = begin
+ ∣ (+ x) - (+ y) ∣ ≡⟨ cong ∣_∣ ([+m]-[+n]≡m⊖n x y) ⟩
+ ∣ x ⊖ y ∣ ≡⟨ ∣m⊖n∣≡∣n⊖m∣ x y ⟩
+ ∣ y ⊖ x ∣ ≡⟨ cong ∣_∣ (sym ([+m]-[+n]≡m⊖n y x)) ⟩
+ ∣ (+ y) - (+ x) ∣ ∎
------------------------------------------------------------------------
-- Properties of _+_
+-comm : Commutative _+_
-+-comm -[1+ a ] -[1+ b ] rewrite ℕₚ.+-comm a b = refl
-+-comm (+ a ) (+ b ) rewrite ℕₚ.+-comm a b = refl
++-comm -[1+ a ] -[1+ b ] = cong (-[1+_] ∘ suc) (ℕₚ.+-comm a b)
++-comm (+ a ) (+ b ) = cong +_ (ℕₚ.+-comm a b)
+-comm -[1+ _ ] (+ _ ) = refl
+-comm (+ _ ) -[1+ _ ] = refl
@@ -184,31 +239,31 @@ sign-⊖-≰ = sign-⊖-< ∘ ℕₚ.≰⇒>
+-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
-
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ˡ-⊖-+-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ʳ-⊖-+-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
+distribʳ-⊖-+-pos a b c = begin
+ + a + (b ⊖ c) ≡⟨ +-comm (+ a) (b ⊖ c) ⟩
+ (b ⊖ c) + + a ≡⟨ distribˡ-⊖-+-pos a b c ⟩
+ b ℕ+ a ⊖ c ≡⟨ cong (_⊖ c) (ℕₚ.+-comm b a) ⟩
+ a ℕ+ b ⊖ c ∎
+
+distribʳ-⊖-+-neg : ∀ a b c → -[1+ a ] + (b ⊖ c) ≡ b ⊖ (suc a ℕ+ c)
+distribʳ-⊖-+-neg a b c = begin
+ -[1+ a ] + (b ⊖ c) ≡⟨ +-comm -[1+ a ] (b ⊖ c) ⟩
+ (b ⊖ c) + -[1+ a ] ≡⟨ distribˡ-⊖-+-neg a b c ⟩
+ b ⊖ suc (c ℕ+ a) ≡⟨ cong (λ x → b ⊖ suc x) (ℕₚ.+-comm c a) ⟩
+ b ⊖ suc (a ℕ+ c) ∎
+-assoc : Associative _+_
+-assoc (+ zero) y z rewrite +-identityˡ y | +-identityˡ (y + z) = refl
@@ -241,31 +296,31 @@ distribʳ-⊖-+-pos a b c
rewrite ℕₚ.+-assoc (suc a) (suc b) (suc c)
= refl
-inverseˡ : LeftInverse (+ 0) -_ _+_
-inverseˡ -[1+ n ] = n⊖n≡0 n
-inverseˡ (+ zero) = refl
-inverseˡ (+ suc n) = n⊖n≡0 n
++-inverseˡ : LeftInverse (+ 0) -_ _+_
++-inverseˡ -[1+ n ] = n⊖n≡0 n
++-inverseˡ (+ zero) = refl
++-inverseˡ (+ suc n) = n⊖n≡0 n
-inverseʳ : RightInverse (+ 0) -_ _+_
-inverseʳ = comm+invˡ⇒invʳ +-comm inverseˡ
++-inverseʳ : RightInverse (+ 0) -_ _+_
++-inverseʳ = comm+invˡ⇒invʳ +-comm +-inverseˡ
+-inverse : Inverse (+ 0) -_ _+_
-+-inverse = inverseˡ , inverseʳ
++-inverse = +-inverseˡ , +-inverseʳ
-+-isSemigroup : IsSemigroup _≡_ _+_
++-isSemigroup : IsSemigroup _+_
+-isSemigroup = record
{ isEquivalence = isEquivalence
; assoc = +-assoc
; ∙-cong = cong₂ _+_
}
-+-0-isMonoid : IsMonoid _≡_ _+_ (+ 0)
++-0-isMonoid : IsMonoid _+_ (+ 0)
+-0-isMonoid = record
{ isSemigroup = +-isSemigroup
; identity = +-identity
}
-+-0-isCommutativeMonoid : IsCommutativeMonoid _≡_ _+_ (+ 0)
++-0-isCommutativeMonoid : IsCommutativeMonoid _+_ (+ 0)
+-0-isCommutativeMonoid = record
{ isSemigroup = +-isSemigroup
; identityˡ = +-identityˡ
@@ -281,14 +336,14 @@ inverseʳ = comm+invˡ⇒invʳ +-comm inverseˡ
; isCommutativeMonoid = +-0-isCommutativeMonoid
}
-+-0-isGroup : IsGroup _≡_ _+_ (+ 0) (-_)
++-0-isGroup : IsGroup _+_ (+ 0) (-_)
+-0-isGroup = record
{ isMonoid = +-0-isMonoid
; inverse = +-inverse
; ⁻¹-cong = cong (-_)
}
-+-isAbelianGroup : IsAbelianGroup _≡_ _+_ (+ 0) (-_)
++-isAbelianGroup : IsAbelianGroup _+_ (+ 0) (-_)
+-isAbelianGroup = record
{ isGroup = +-0-isGroup
; comm = +-comm
@@ -304,9 +359,6 @@ inverseʳ = comm+invˡ⇒invʳ +-comm inverseˡ
; isAbelianGroup = +-isAbelianGroup
}
-open Algebra.Properties.AbelianGroup +-0-abelianGroup
- using () renaming (⁻¹-involutive to -‿involutive)
-
-- Other properties of _+_
n≢1+n : ∀ {n} → n ≢ sucℤ n
@@ -323,9 +375,9 @@ 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) =
+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)
@@ -341,6 +393,14 @@ neg-distrib-+ -[1+ m ] (+ n) =
(+ m) + (+ n) ≡⟨ sym (cong₂ _+_ (+◃n≡+n m) (+◃n≡+n n)) ⟩
(Sign.+ ◃ m) + (Sign.+ ◃ n) ∎
++-minus-telescope : ∀ x y z → (x - y) + (y - z) ≡ x - z
++-minus-telescope x y z = begin
+ (x - y) + (y - z) ≡⟨ +-assoc x (- y) (y - z) ⟩
+ x + (- y + (y - z)) ≡⟨ cong (λ v → x + v) (sym (+-assoc (- y) y _)) ⟩
+ x + ((- y + y) - z) ≡⟨ sym (+-assoc x (- y + y) (- z)) ⟩
+ x + (- y + y) - z ≡⟨ cong (λ a → x + a - z) (+-inverseˡ y) ⟩
+ x + (+ 0) - z ≡⟨ cong (_- z) (+-identityʳ x) ⟩
+ x - z ∎
------------------------------------------------------------------------
-- Properties of _*_
@@ -379,11 +439,10 @@ private
:= 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 (+ zero) z rewrite ℕₚ.*-zeroʳ ∣ x ∣ = refl
*-assoc x y (+ zero) rewrite
ℕₚ.*-zeroʳ ∣ y ∣
| ℕₚ.*-zeroʳ ∣ x ∣
@@ -398,48 +457,13 @@ private
*-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₂ _*_
- }
-
-*-1-isMonoid : IsMonoid _≡_ _*_ (+ 1)
-*-1-isMonoid = record
- { isSemigroup = *-isSemigroup
- ; identity = *-identity
- }
-
-*-1-isCommutativeMonoid : IsCommutativeMonoid _≡_ _*_ (+ 1)
-*-1-isCommutativeMonoid = record
- { isSemigroup = *-isSemigroup
- ; identityˡ = *-identityˡ
- ; comm = *-comm
- }
-
-*-1-commutativeMonoid : CommutativeMonoid _ _
-*-1-commutativeMonoid = record
- { Carrier = ℤ
- ; _≈_ = _≡_
- ; _∙_ = _*_
- ; ε = + 1
- ; isCommutativeMonoid = *-1-isCommutativeMonoid
- }
-
-------------------------------------------------------------------------
--- The integers form a commutative ring
-
--- Distributivity
-
private
-- 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)
+ rewrite +-cancelˡ-⊖ a (b ℕ* suc a) (c ℕ* suc a)
| ⊖-swap (b ℕ* suc a) (c ℕ* suc a)
with b ℕ≤? c
... | yes b≤c
@@ -453,115 +477,133 @@ private
| ∣⊖∣-≰ b≰c
| +◃n≡+n ((b ∸ c) ℕ* suc a)
| ⊖-≰ (b≰c ∘ ℕₚ.*-cancelʳ-≤ b c a)
- | -‿involutive (+ (b ℕ* suc a ∸ c ℕ* suc a))
+ | neg-involutive (+ (b ℕ* suc a ∸ c ℕ* suc a))
| ℕₚ.*-distribʳ-∸ (suc a) b c
= refl
-distribʳ : _*_ DistributesOverʳ _+_
-
-distribʳ (+ zero) y z
+*-distribʳ-+ : _*_ DistributesOverʳ _+_
+*-distribʳ-+ (+ zero) y z
rewrite ℕₚ.*-zeroʳ ∣ y ∣
| ℕₚ.*-zeroʳ ∣ z ∣
| ℕₚ.*-zeroʳ ∣ y + z ∣
= refl
-
-distribʳ x (+ zero) z
+*-distribʳ-+ x (+ zero) z
rewrite +-identityˡ z
| +-identityˡ (sign z 𝕊* sign x ◃ ∣ z ∣ ℕ* ∣ x ∣)
= refl
-
-distribʳ x y (+ zero)
+*-distribʳ-+ x y (+ zero)
rewrite +-identityʳ y
| +-identityʳ (sign y 𝕊* sign x ◃ ∣ y ∣ ℕ* ∣ x ∣)
= refl
-
-distribʳ -[1+ a ] -[1+ b ] -[1+ c ] = cong (+_) $
+*-distribʳ-+ -[1+ a ] -[1+ b ] -[1+ c ] = cong (+_) $
solve 3 (λ a b c → (con 2 :+ b :+ c) :* (con 1 :+ a)
:= (con 1 :+ b) :* (con 1 :+ a) :+
(con 1 :+ c) :* (con 1 :+ a))
refl a b c
-
-distribʳ (+ suc a) (+ suc b) (+ suc c) = cong (+_) $
+*-distribʳ-+ (+ suc a) (+ suc b) (+ suc c) = cong (+_) $
solve 3 (λ a b c → (con 1 :+ b :+ (con 1 :+ c)) :* (con 1 :+ a)
:= (con 1 :+ b) :* (con 1 :+ a) :+
(con 1 :+ c) :* (con 1 :+ a))
refl a b c
-
-distribʳ -[1+ a ] (+ suc b) (+ suc c) = cong -[1+_] $
+*-distribʳ-+ -[1+ a ] (+ suc b) (+ suc c) = cong -[1+_] $
solve 3 (λ a b c → a :+ (b :+ (con 1 :+ c)) :* (con 1 :+ a)
:= (con 1 :+ b) :* (con 1 :+ a) :+
(a :+ c :* (con 1 :+ a)))
refl a b c
-
-distribʳ (+ suc a) -[1+ b ] -[1+ c ] = cong -[1+_] $
+*-distribʳ-+ (+ suc a) -[1+ b ] -[1+ c ] = cong -[1+_] $
solve 3 (λ a b c → a :+ (con 1 :+ a :+ (b :+ c) :* (con 1 :+ a))
:= (con 1 :+ b) :* (con 1 :+ a) :+
(a :+ c :* (con 1 :+ a)))
refl a b c
-
-distribʳ -[1+ a ] -[1+ b ] (+ suc c) = distrib-lemma a b c
-
-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
+*-distribʳ-+ -[1+ a ] -[1+ b ] (+ suc c) = distrib-lemma a b c
+*-distribʳ-+ -[1+ a ] (+ suc b) -[1+ c ] = distrib-lemma a c b
+*-distribʳ-+ (+ suc a) -[1+ b ] (+ suc c) with b ℕ≤? c
... | yes b≤c
- rewrite ⊖-≥ b≤c
+ rewrite +-cancelˡ-⊖ a (c ℕ* suc a) (b ℕ* suc a)
+ | ⊖-≥ b≤c
| +-comm (- (+ (a ℕ+ b ℕ* suc a))) (+ (a ℕ+ c ℕ* 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
+ rewrite +-cancelˡ-⊖ a (c ℕ* suc a) (b ℕ* suc a)
+ | 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
+*-distribʳ-+ (+ suc c) (+ suc a) -[1+ b ] with b ℕ≤? a
... | yes b≤a
- rewrite ⊖-≥ b≤a
+ rewrite +-cancelˡ-⊖ c (a ℕ* suc c) (b ℕ* suc c)
+ | ⊖-≥ b≤a
| ⊖-≥ (ℕₚ.*-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
+ rewrite +-cancelˡ-⊖ c (a ℕ* suc c) (b ℕ* suc c)
+ | sign-⊖-≰ b≰a
| ∣⊖∣-≰ b≰a
| ⊖-≰ (b≰a ∘ ℕₚ.*-cancelʳ-≤ b a c)
| -◃n≡-n ((b ∸ a) ℕ* suc c)
| ℕₚ.*-distribʳ-∸ (suc c) b a
= refl
-isCommutativeSemiring : IsCommutativeSemiring _≡_ _+_ _*_ (+ 0) (+ 1)
-isCommutativeSemiring = record
+*-isSemigroup : IsSemigroup _*_
+*-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = *-assoc
+ ; ∙-cong = cong₂ _*_
+ }
+
+*-1-isMonoid : IsMonoid _*_ (+ 1)
+*-1-isMonoid = record
+ { isSemigroup = *-isSemigroup
+ ; identity = *-identity
+ }
+
+*-1-isCommutativeMonoid : IsCommutativeMonoid _*_ (+ 1)
+*-1-isCommutativeMonoid = record
+ { isSemigroup = *-isSemigroup
+ ; identityˡ = *-identityˡ
+ ; comm = *-comm
+ }
+
+*-1-commutativeMonoid : CommutativeMonoid _ _
+*-1-commutativeMonoid = record
+ { Carrier = ℤ
+ ; _≈_ = _≡_
+ ; _∙_ = _*_
+ ; ε = + 1
+ ; isCommutativeMonoid = *-1-isCommutativeMonoid
+ }
+
++-*-isCommutativeSemiring : IsCommutativeSemiring _+_ _*_ (+ 0) (+ 1)
++-*-isCommutativeSemiring = record
{ +-isCommutativeMonoid = +-0-isCommutativeMonoid
; *-isCommutativeMonoid = *-1-isCommutativeMonoid
- ; distribʳ = distribʳ
- ; zeroˡ = λ _ → refl
+ ; distribʳ = *-distribʳ-+
+ ; zeroˡ = *-zeroˡ
}
-+-*-isRing : IsRing _≡_ _+_ _*_ -_ (+ 0) (+ 1)
++-*-isRing : IsRing _+_ _*_ -_ (+ 0) (+ 1)
+-*-isRing = record
{ +-isAbelianGroup = +-isAbelianGroup
; *-isMonoid = *-1-isMonoid
; distrib = IsCommutativeSemiring.distrib
- isCommutativeSemiring
+ +-*-isCommutativeSemiring
}
-+-*-isCommutativeRing : IsCommutativeRing _≡_ _+_ _*_ -_ (+ 0) (+ 1)
++-*-isCommutativeRing : IsCommutativeRing _+_ _*_ -_ (+ 0) (+ 1)
+-*-isCommutativeRing = record
{ isRing = +-*-isRing
; *-comm = *-comm
}
-commutativeRing : CommutativeRing _ _
-commutativeRing = record
++-*-commutativeRing : CommutativeRing _ _
++-*-commutativeRing = record
{ Carrier = ℤ
; _≈_ = _≡_
; _+_ = _+_
@@ -572,28 +614,40 @@ commutativeRing = record
; isCommutativeRing = +-*-isCommutativeRing
}
-import Algebra.RingSolver.Simple as Solver
-import Algebra.RingSolver.AlmostCommutativeRing as ACR
-module RingSolver =
- Solver (ACR.fromCommutativeRing commutativeRing) _≟_
-
-- Other properties of _*_
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 ≢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
+pos-distrib-* : ∀ x y → (+ x) * (+ y) ≡ + (x ℕ* y)
+pos-distrib-* zero y = refl
+pos-distrib-* (suc x) zero = pos-distrib-* x zero
+pos-distrib-* (suc x) (suc y) = refl
+
+◃-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)))
+
+*-cancelʳ-≡ : ∀ i j k → k ≢ + 0 → i * k ≡ j * k → i ≡ j
+*-cancelʳ-≡ i j k ≢0 eq with signAbs k
+*-cancelʳ-≡ i j .(+ 0) ≢0 eq | s ◂ zero = contradiction refl ≢0
+*-cancelʳ-≡ 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ʳ-≡ ∣ i ∣ ∣ j ∣ $ abs-cong eq
where
sign-i≡sign-j : ∀ i j →
- sign i 𝕊* s ◃ ∣ i ∣ ℕ* suc n ≡
- sign j 𝕊* 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
@@ -613,29 +667,29 @@ 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 =
- 𝕊ₚ.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ʳ-≤ (suc n) (suc m) o (s≤s n≤m)))
-cancel-*-+-right-≤ -[1+ _ ] (+ _) _ _ = -≤+
-cancel-*-+-right-≤ (+ 0) -[1+ _ ] _ ()
-cancel-*-+-right-≤ (+ suc _) -[1+ _ ] _ ()
-cancel-*-+-right-≤ (+ 0) (+ 0) _ _ = +≤+ z≤n
-cancel-*-+-right-≤ (+ 0) (+ suc _) _ _ = +≤+ z≤n
-cancel-*-+-right-≤ (+ suc _) (+ 0) _ (+≤+ ())
-cancel-*-+-right-≤ (+ suc m) (+ suc n) o (+≤+ m≤n) =
+ 𝕊ₚ.*-cancelʳ-≡ s₁ s₂ (sign-cong eq)
+
+*-cancelʳ-≤-pos : ∀ m n o → m * + suc o ≤ n * + suc o → m ≤ n
+*-cancelʳ-≤-pos (-[1+ m ]) (-[1+ n ]) o (-≤- n≤m) =
+ -≤- (ℕₚ.≤-pred (ℕₚ.*-cancelʳ-≤ (suc n) (suc m) o (s≤s n≤m)))
+*-cancelʳ-≤-pos -[1+ _ ] (+ _) _ _ = -≤+
+*-cancelʳ-≤-pos (+ 0) -[1+ _ ] _ ()
+*-cancelʳ-≤-pos (+ suc _) -[1+ _ ] _ ()
+*-cancelʳ-≤-pos (+ 0) (+ 0) _ _ = +≤+ z≤n
+*-cancelʳ-≤-pos (+ 0) (+ suc _) _ _ = +≤+ z≤n
+*-cancelʳ-≤-pos (+ suc _) (+ 0) _ (+≤+ ())
+*-cancelʳ-≤-pos (+ suc m) (+ suc n) o (+≤+ m≤n) =
+≤+ (ℕₚ.*-cancelʳ-≤ (suc m) (suc n) o m≤n)
-*-+-right-mono : ∀ n → (_* + suc n) Preserves _≤_ ⟶ _≤_
-*-+-right-mono _ (-≤+ {n = 0}) = -≤+
-*-+-right-mono _ (-≤+ {n = suc _}) = -≤+
-*-+-right-mono x (-≤- n≤m) =
- -≤- (≤-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) =
+*-monoʳ-≤-pos : ∀ n → (_* + suc n) Preserves _≤_ ⟶ _≤_
+*-monoʳ-≤-pos _ (-≤+ {n = 0}) = -≤+
+*-monoʳ-≤-pos _ (-≤+ {n = suc _}) = -≤+
+*-monoʳ-≤-pos x (-≤- n≤m) =
+ -≤- (ℕₚ.≤-pred (ℕₚ.*-mono-≤ (s≤s n≤m) (ℕₚ.≤-refl {x = suc x})))
+*-monoʳ-≤-pos _ (+≤+ {m = 0} {n = 0} m≤n) = +≤+ m≤n
+*-monoʳ-≤-pos _ (+≤+ {m = 0} {n = suc _} m≤n) = +≤+ z≤n
+*-monoʳ-≤-pos _ (+≤+ {m = suc _} {n = 0} ())
+*-monoʳ-≤-pos x (+≤+ {m = suc _} {n = suc _} m≤n) =
+≤+ ((ℕₚ.*-mono-≤ m≤n (ℕₚ.≤-refl {x = suc x})))
-1*n≡-n : ∀ n → -[1+ 0 ] * n ≡ - n
@@ -643,21 +697,15 @@ cancel-*-+-right-≤ (+ suc m) (+ suc n) o (+≤+ m≤n) =
-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 _≤_
+drop‿+≤+ : ∀ {m n} → + m ≤ + n → m ℕ≤ n
+drop‿+≤+ (+≤+ m≤n) = m≤n
+
+drop‿-≤- : ∀ {m n} → -[1+ m ] ≤ -[1+ n ] → n ℕ≤ m
+drop‿-≤- (-≤- n≤m) = n≤m
+
≤-reflexive : _≡_ ⇒ _≤_
≤-reflexive { -[1+ n ]} refl = -≤- ℕₚ.≤-refl
≤-reflexive {+ n} refl = +≤+ ℕₚ.≤-refl
@@ -686,6 +734,13 @@ cancel-*-+-right-≤ (+ suc m) (+ suc n) o (+≤+ m≤n) =
... | inj₁ m≤n = inj₁ (+≤+ m≤n)
... | inj₂ n≤m = inj₂ (+≤+ n≤m)
+infix 4 _≤?_
+_≤?_ : Decidable _≤_
+-[1+ m ] ≤? -[1+ n ] = Dec.map′ -≤- drop‿-≤- (ℕ._≤?_ n m)
+-[1+ m ] ≤? + n = yes -≤+
++ m ≤? -[1+ n ] = no λ ()
++ m ≤? + n = Dec.map′ +≤+ drop‿+≤+ (ℕ._≤?_ m n)
+
≤-isPreorder : IsPreorder _≡_ _≤_
≤-isPreorder = record
{ isEquivalence = isEquivalence
@@ -728,9 +783,6 @@ cancel-*-+-right-≤ (+ suc m) (+ suc n) o (+≤+ m≤n) =
; 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)
@@ -740,6 +792,11 @@ module ≤-Reasoning = POR ≤-poset renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
n≤1+n : ∀ n → n ≤ (+ 1) + n
n≤1+n n = ≤-step ≤-refl
+≤-irrelevance : Irrelevant _≤_
+≤-irrelevance -≤+ -≤+ = refl
+≤-irrelevance (-≤- n≤m₁) (-≤- n≤m₂) = cong -≤- (ℕₚ.≤-irrelevance n≤m₁ n≤m₂)
+≤-irrelevance (+≤+ n≤m₁) (+≤+ n≤m₂) = cong +≤+ (ℕₚ.≤-irrelevance n≤m₁ n≤m₂)
+
------------------------------------------------------------------------
-- Properties _<_
@@ -789,11 +846,11 @@ n≤1+n n = ≤-step ≤-refl
<-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 ∘ 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‿-≤-)
+ tri< (-≤- (ℕₚ.≤-pred m>n)) (m≢n ∘ -[1+-injective) (m≮n ∘ s≤s ∘ drop‿-≤-)
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
<-isStrictTotalOrder = record
@@ -827,3 +884,76 @@ n≮n { -[1+ suc n ]} (-≤- n<n) = contradiction n<n ℕₚ.1+n≰n
≰→> { -[1+ m ]} { -[1+ suc n ]} m≰n with m ℕ≤? n
... | yes m≤n = -≤- m≤n
... | no m≰n' = contradiction (-≤- (ℕₚ.≰⇒> m≰n')) m≰n
+
+<-irrelevance : Irrelevant _<_
+<-irrelevance = ≤-irrelevance
+
+------------------------------------------------------------------------
+-- Modules for reasoning about integer number relations
+
+-- A module for reasoning about the _≤_ relation
+module ≤-Reasoning = POR ≤-poset hiding (_≈⟨_⟩_)
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+inverseˡ = +-inverseˡ
+{-# WARNING_ON_USAGE inverseˡ
+"Warning: inverseˡ was deprecated in v0.15.
+Please use +-inverseˡ instead."
+#-}
+inverseʳ = +-inverseʳ
+{-# WARNING_ON_USAGE inverseʳ
+"Warning: inverseʳ was deprecated in v0.15.
+Please use +-inverseʳ instead."
+#-}
+distribʳ = *-distribʳ-+
+{-# WARNING_ON_USAGE distribʳ
+"Warning: distribʳ was deprecated in v0.15.
+Please use *-distribʳ-+ instead."
+#-}
+isCommutativeSemiring = +-*-isCommutativeSemiring
+{-# WARNING_ON_USAGE isCommutativeSemiring
+"Warning: isCommutativeSemiring was deprecated in v0.15.
+Please use +-*-isCommutativeSemiring instead."
+#-}
+commutativeRing = +-*-commutativeRing
+{-# WARNING_ON_USAGE commutativeRing
+"Warning: commutativeRing was deprecated in v0.15.
+Please use +-*-commutativeRing instead."
+#-}
+*-+-right-mono = *-monoʳ-≤-pos
+{-# WARNING_ON_USAGE *-+-right-mono
+"Warning: *-+-right-mono was deprecated in v0.15.
+Please use *-monoʳ-≤-pos instead."
+#-}
+cancel-*-+-right-≤ = *-cancelʳ-≤-pos
+{-# WARNING_ON_USAGE cancel-*-+-right-≤
+"Warning: cancel-*-+-right-≤ was deprecated in v0.15.
+Please use *-cancelʳ-≤-pos instead."
+#-}
+cancel-*-right = *-cancelʳ-≡
+{-# WARNING_ON_USAGE cancel-*-right
+"Warning: cancel-*-right was deprecated in v0.15.
+Please use *-cancelʳ-≡ instead."
+#-}
+doubleNeg = neg-involutive
+{-# WARNING_ON_USAGE doubleNeg
+"Warning: doubleNeg was deprecated in v0.15.
+Please use neg-involutive instead."
+#-}
+-‿involutive = neg-involutive
+{-# WARNING_ON_USAGE -‿involutive
+"Warning: -‿involutive was deprecated in v0.15.
+Please use neg-involutive instead."
+#-}
++-⊖-left-cancel = +-cancelˡ-⊖
+{-# WARNING_ON_USAGE +-⊖-left-cancel
+"Warning: +-⊖-left-cancel was deprecated in v0.15.
+Please use +-cancelˡ-⊖ instead."
+#-}
diff --git a/src/Data/Integer/Solver.agda b/src/Data/Integer/Solver.agda
new file mode 100644
index 0000000..024f075
--- /dev/null
+++ b/src/Data/Integer/Solver.agda
@@ -0,0 +1,22 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Automatic solvers for equations over integers
+------------------------------------------------------------------------
+
+-- See README.Integer for examples of how to use this solver
+
+module Data.Integer.Solver where
+
+import Algebra.Solver.Ring.Simple as Solver
+import Algebra.Solver.Ring.AlmostCommutativeRing as ACR
+open import Data.Integer using (_≟_)
+open import Data.Integer.Properties using (+-*-commutativeRing)
+
+------------------------------------------------------------------------
+-- A module for automatically solving propositional equivalences
+-- containing _+_ and _*_
+
+-- A module for automatically solving propositional equivalences
+module +-*-Solver =
+ Solver (ACR.fromCommutativeRing +-*-commutativeRing) _≟_
diff --git a/src/Data/List.agda b/src/Data/List.agda
index 211b7f2..2de19a0 100644
--- a/src/Data/List.agda
+++ b/src/Data/List.agda
@@ -6,87 +6,7 @@
module Data.List where
-open import Data.Nat
-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 Algebra
-import Relation.Binary.PropositionalEquality as PropEq
-import Algebra.FunctionProperties as FunProp
-
------------------------------------------------------------------------
-- Types and basic operations
open import Data.List.Base public
-
-------------------------------------------------------------------------
--- List monoid
-
-monoid : ∀ {ℓ} → Set ℓ → Monoid _ _
-monoid A = record
- { Carrier = List A
- ; _≈_ = _≡_
- ; _∙_ = _++_
- ; ε = []
- ; isMonoid = record
- { isSemigroup = record
- { isEquivalence = PropEq.isEquivalence
- ; assoc = assoc
- ; ∙-cong = cong₂ _++_
- }
- ; identity = ((λ _ → refl) , identity)
- }
- }
- where
- open PropEq
- open FunProp _≡_
-
- identity : RightIdentity [] _++_
- identity [] = refl
- identity (x ∷ xs) = cong (_∷_ x) (identity xs)
-
- assoc : Associative _++_
- assoc [] ys zs = refl
- assoc (x ∷ xs) ys zs = cong (_∷_ x) (assoc xs ys zs)
-
-------------------------------------------------------------------------
--- List monad
-
-open import Category.Monad
-
-monad : ∀ {ℓ} → RawMonad (List {ℓ})
-monad = record
- { return = λ x → x ∷ []
- ; _>>=_ = λ xs f → concat (map f xs)
- }
-
-monadZero : ∀ {ℓ} → RawMonadZero (List {ℓ})
-monadZero = record
- { monad = monad
- ; ∅ = []
- }
-
-monadPlus : ∀ {ℓ} → RawMonadPlus (List {ℓ})
-monadPlus = record
- { monadZero = monadZero
- ; _∣_ = _++_
- }
-
-------------------------------------------------------------------------
--- Monadic functions
-
-private
- module Monadic {m} {M : Set m → Set m} (Mon : RawMonad M) where
-
- open RawMonad Mon
-
- sequence : ∀ {A} → List (M A) → M (List A)
- sequence [] = return []
- sequence (x ∷ xs) = _∷_ <$> x ⊛ sequence xs
-
- mapM : ∀ {a} {A : Set a} {B} → (A → M B) → List A → M (List B)
- mapM f = sequence ∘ map f
-
-open Monadic public
diff --git a/src/Data/List/All.agda b/src/Data/List/All.agda
index a1caee8..88a5cfe 100644
--- a/src/Data/List/All.agda
+++ b/src/Data/List/All.agda
@@ -6,16 +6,18 @@
module Data.List.All where
-open import Data.List.Base as List hiding (map; tabulate; all)
+open import Data.List.Base as List using (List; []; _∷_)
open import Data.List.Any as Any using (here; there)
-open import Data.List.Any.Membership.Propositional using (_∈_)
+open import Data.List.Membership.Propositional using (_∈_)
+open import Data.Product as Prod using (_,_)
open import Function
open import Level
open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
-open import Relation.Unary using (Decidable) renaming (_⊆_ to _⋐_)
-open import Relation.Binary.PropositionalEquality
+open import Relation.Unary hiding (_∈_)
+open import Relation.Binary.PropositionalEquality as P
+------------------------------------------------------------------------
-- All P xs means that all elements in xs satisfy P.
infixr 5 _∷_
@@ -25,6 +27,9 @@ data All {a p} {A : Set a}
[] : All P []
_∷_ : ∀ {x xs} (px : P x) (pxs : All P xs) → All P (x ∷ xs)
+------------------------------------------------------------------------
+-- Operations on All
+
head : ∀ {a p} {A : Set a} {P : A → Set p} {x xs} →
All P (x ∷ xs) → P x
head (px ∷ pxs) = px
@@ -34,7 +39,7 @@ tail : ∀ {a p} {A : Set a} {P : A → Set p} {x xs} →
tail (px ∷ pxs) = pxs
lookup : ∀ {a p} {A : Set a} {P : A → Set p} {xs : List A} →
- All P xs → (∀ {x : A} → x ∈ xs → P x)
+ All P xs → (∀ {x} → x ∈ xs → P x)
lookup [] ()
lookup (px ∷ pxs) (here refl) = px
lookup (px ∷ pxs) (there x∈xs) = lookup pxs x∈xs
@@ -45,13 +50,39 @@ tabulate {xs = []} hyp = []
tabulate {xs = x ∷ xs} hyp = hyp (here refl) ∷ tabulate (hyp ∘ there)
map : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} →
- P ⋐ Q → All P ⋐ All Q
+ P ⊆ Q → All P ⊆ All Q
map g [] = []
map g (px ∷ pxs) = g px ∷ map g pxs
-all : ∀ {a p} {A : Set a} {P : A → Set p} →
- Decidable P → Decidable (All P)
-all p [] = yes []
-all p (x ∷ xs) with p x
-all p (x ∷ xs) | yes px = Dec.map′ (_∷_ px) tail (all p xs)
-all p (x ∷ xs) | no ¬px = no (¬px ∘ head)
+zip : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} →
+ All P ∩ All Q ⊆ All (P ∩ Q)
+zip ([] , []) = []
+zip (px ∷ pxs , qx ∷ qxs) = (px , qx) ∷ zip (pxs , qxs)
+
+unzip : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} →
+ All (P ∩ Q) ⊆ All P ∩ All Q
+unzip [] = [] , []
+unzip (pqx ∷ pqxs) = Prod.zip _∷_ _∷_ pqx (unzip pqxs)
+
+------------------------------------------------------------------------
+-- Properties of predicates preserved by All
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ all : Decidable P → Decidable (All P)
+ all p [] = yes []
+ all p (x ∷ xs) with p x
+ ... | yes px = Dec.map′ (px ∷_) tail (all p xs)
+ ... | no ¬px = no (¬px ∘ head)
+
+ universal : Universal P → Universal (All P)
+ universal u [] = []
+ universal u (x ∷ xs) = u x ∷ universal u xs
+
+ irrelevant : Irrelevant P → Irrelevant (All P)
+ irrelevant irr [] [] = P.refl
+ irrelevant irr (px₁ ∷ pxs₁) (px₂ ∷ pxs₂) =
+ P.cong₂ _∷_ (irr px₁ px₂) (irrelevant irr pxs₁ pxs₂)
+
+ satisfiable : Satisfiable (All P)
+ satisfiable = [] , []
diff --git a/src/Data/List/All/Properties.agda b/src/Data/List/All/Properties.agda
index 9830f36..ff0cb3a 100644
--- a/src/Data/List/All/Properties.agda
+++ b/src/Data/List/All/Properties.agda
@@ -11,31 +11,25 @@ open import Data.Bool.Properties
open import Data.Empty
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.Membership.Propositional
open import Data.List.All as All using (All; []; _∷_)
open import Data.List.Any using (Any; here; there)
+open import Data.List.Relation.Pointwise using (Pointwise; []; _∷_)
+open import Data.List.Relation.Subset.Propositional using (_⊆_)
+open import Data.Maybe as Maybe using (Maybe; just; nothing)
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)
-open import Function.Inverse using (_↔_)
-open import Function.Surjection using (_↠_)
+open import Function.Equivalence using (_⇔_; equivalence; Equivalence)
+open import Function.Inverse using (_↔_; inverse)
+open import Function.Surjection using (_↠_; surjection)
open import Relation.Binary.PropositionalEquality as P using (_≡_)
open import Relation.Nullary
open import Relation.Unary
using (Decidable; Universal) renaming (_⊆_ to _⋐_)
------------------------------------------------------------------------
--- 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
@@ -60,15 +54,8 @@ module _ {a p} {A : Set a} {P : A → Set p} where
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
- }
- }
+ ¬Any↠All¬ = surjection (¬Any⇒All¬ _) All¬⇒¬Any 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)
@@ -85,66 +72,49 @@ module _ {a p} {A : Set a} {P : A → Set p} where
}
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 _)
- }
+ Any¬⇔¬All dec = equivalence Any¬→¬All (¬All⇒Any¬ dec _)
where
-
-- If equality of functions were extensional, then the logical
-- equivalence could be strengthened to a surjection.
-
to∘from : P.Extensionality _ _ →
∀ {xs} (¬∀ : ¬ All P xs) → Any¬→¬All (¬All⇒Any¬ dec xs ¬∀) ≡ ¬∀
to∘from ext ¬∀ = ext (⊥-elim ∘ ¬∀)
------------------------------------------------------------------------
--- Lemmas relating All to ⊤
-
-All-all : ∀ {a} {A : Set a} (p : A → Bool) {xs} →
- All (T ∘ p) xs → T (all p xs)
-All-all p [] = _
-All-all p (px ∷ pxs) = Equivalence.from T-∧ ⟨$⟩ (px , All-all p pxs)
-
-all-All : ∀ {a} {A : Set a} (p : A → Bool) xs →
- T (all p xs) → All (T ∘ p) xs
-all-All p [] _ = []
-all-All p (x ∷ xs) px∷xs with Equivalence.to (T-∧ {p x}) ⟨$⟩ px∷xs
-all-All p (x ∷ xs) px∷xs | (px , pxs) = px ∷ all-All p xs pxs
-
+-- Introduction (⁺) and elimination (⁻) rules for list operations
------------------------------------------------------------------------
--- All is anti-monotone.
+-- map
-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)
+module _ {a b p} {A : Set a} {B : Set b} {P : B → Set p} {f : A → B} where
-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 _
+ map⁺ : ∀ {xs} → All (P ∘ f) xs → All P (map f xs)
+ map⁺ [] = []
+ map⁺ (p ∷ ps) = p ∷ map⁺ ps
-------------------------------------------------------------------------
--- Introduction (⁺) and elimination (⁻) rules for various list functions
-------------------------------------------------------------------------
--- map
+ map⁻ : ∀ {xs} → All P (map f xs) → All (P ∘ f) xs
+ map⁻ {xs = []} [] = []
+ map⁻ {xs = _ ∷ _} (p ∷ ps) = p ∷ map⁻ ps
-module _{a b} {A : Set a} {B : Set b} where
+-- A variant of All.map.
- 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
+module _ {a b p q} {A : Set a} {B : Set b} {f : A → B}
+ {P : A → Set p} {Q : B → Set q} where
- 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
+ gmap : P ⋐ Q ∘ f → All P ⋐ All Q ∘ map f
+ gmap g = map⁺ ∘ All.map g
- -- A variant of All.map.
+------------------------------------------------------------------------
+-- mapMaybe
- 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
+module _ {a b p} {A : Set a} {B : Set b}
+ (P : B → Set p) {f : A → Maybe B} where
+
+ mapMaybe⁺ : ∀ {xs} → All (Maybe.All P) (map f xs) → All P (mapMaybe f xs)
+ mapMaybe⁺ {[]} [] = []
+ mapMaybe⁺ {x ∷ xs} (px ∷ pxs) with f x
+ ... | nothing = mapMaybe⁺ pxs
+ ... | just v with px
+ ... | just pv = pv ∷ mapMaybe⁺ pxs
------------------------------------------------------------------------
-- _++_
@@ -168,16 +138,8 @@ module _ {a p} {A : Set a} {P : A → Set p} where
++⁻ (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
- }
- }
+ ++↔ {xs} = inverse (uncurry ++⁺) (++⁻ xs) ++⁻∘++⁺ (++⁺∘++⁻ xs)
where
-
++⁺∘++⁻ : ∀ xs {ys} (p : All P (xs ++ ys)) →
uncurry′ ++⁺ (++⁻ xs p) ≡ p
++⁺∘++⁻ [] p = P.refl
@@ -249,3 +211,144 @@ module _ {a p} {A : Set a} {P : A → Set p} where
tabulate⁻ {zero} pf ()
tabulate⁻ {suc n} (px ∷ _) fzero = px
tabulate⁻ {suc n} (_ ∷ pf) (fsuc i) = tabulate⁻ pf i
+
+------------------------------------------------------------------------
+-- filter
+
+module _ {a p} {A : Set a} {P : A → Set p} (P? : Decidable P) where
+
+ filter⁺₁ : ∀ xs → All P (filter P? xs)
+ filter⁺₁ [] = []
+ filter⁺₁ (x ∷ xs) with P? x
+ ... | yes Px = Px ∷ filter⁺₁ xs
+ ... | no _ = filter⁺₁ xs
+
+ filter⁺₂ : ∀ {q} {Q : A → Set q} {xs} →
+ All Q xs → All Q (filter P? xs)
+ filter⁺₂ {xs = _} [] = []
+ filter⁺₂ {xs = x ∷ _} (Qx ∷ Qxs) with P? x
+ ... | no _ = filter⁺₂ Qxs
+ ... | yes _ = Qx ∷ filter⁺₂ Qxs
+
+------------------------------------------------------------------------
+-- zipWith
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
+
+ zipWith⁺ : ∀ {p} (P : C → Set p) (f : A → B → C) {xs ys} →
+ Pointwise (λ x y → P (f x y)) xs ys →
+ All P (zipWith f xs ys)
+ zipWith⁺ P f [] = []
+ zipWith⁺ P f (Pfxy ∷ Pfxsys) = Pfxy ∷ zipWith⁺ P f Pfxsys
+
+------------------------------------------------------------------------
+-- Operations for constructing lists
+------------------------------------------------------------------------
+-- singleton
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ singleton⁻ : ∀ {x} → All P [ x ] → P x
+ singleton⁻ (px ∷ []) = px
+
+------------------------------------------------------------------------
+-- fromMaybe
+
+ fromMaybe⁺ : ∀ {mx} → Maybe.All P mx → All P (fromMaybe mx)
+ fromMaybe⁺ (just px) = px ∷ []
+ fromMaybe⁺ nothing = []
+
+ fromMaybe⁻ : ∀ mx → All P (fromMaybe mx) → Maybe.All P mx
+ fromMaybe⁻ (just x) (px ∷ []) = just px
+ fromMaybe⁻ nothing p = nothing
+
+------------------------------------------------------------------------
+-- replicate
+
+ replicate⁺ : ∀ n {x} → P x → All P (replicate n x)
+ replicate⁺ zero px = []
+ replicate⁺ (suc n) px = px ∷ replicate⁺ n px
+
+ replicate⁻ : ∀ {n x} → All P (replicate (suc n) x) → P x
+ replicate⁻ (px ∷ _) = px
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+------------------------------------------------------------------------
+-- inits
+
+ inits⁺ : ∀ {xs} → All P xs → All (All P) (inits xs)
+ inits⁺ [] = [] ∷ []
+ inits⁺ (px ∷ pxs) = [] ∷ gmap (px ∷_) (inits⁺ pxs)
+
+ inits⁻ : ∀ xs → All (All P) (inits xs) → All P xs
+ inits⁻ [] pxs = []
+ inits⁻ (x ∷ []) ([] ∷ p[x] ∷ []) = p[x]
+ inits⁻ (x ∷ xs@(_ ∷ _)) ([] ∷ pxs@(p[x] ∷ _)) =
+ singleton⁻ p[x] ∷ inits⁻ xs (All.map (drop⁺ 1) (map⁻ pxs))
+
+------------------------------------------------------------------------
+-- tails
+
+ tails⁺ : ∀ {xs} → All P xs → All (All P) (tails xs)
+ tails⁺ [] = [] ∷ []
+ tails⁺ pxxs@(_ ∷ pxs) = pxxs ∷ tails⁺ pxs
+
+ tails⁻ : ∀ xs → All (All P) (tails xs) → All P xs
+ tails⁻ [] pxs = []
+ tails⁻ (x ∷ xs) (pxxs ∷ _) = pxxs
+
+------------------------------------------------------------------------
+-- all
+
+module _ {a} {A : Set a} (p : A → Bool) where
+
+ all⁺ : ∀ xs → T (all p xs) → All (T ∘ p) xs
+ all⁺ [] _ = []
+ all⁺ (x ∷ xs) px∷xs with Equivalence.to (T-∧ {p x}) ⟨$⟩ px∷xs
+ ... | (px , pxs) = px ∷ all⁺ xs pxs
+
+ all⁻ : ∀ {xs} → All (T ∘ p) xs → T (all p xs)
+ all⁻ [] = _
+ all⁻ (px ∷ pxs) = Equivalence.from T-∧ ⟨$⟩ (px , all⁻ 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-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⁻ p ∘ anti-mono xs⊆ys ∘ all⁺ p _
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.16
+
+All-all = all⁻
+{-# WARNING_ON_USAGE All-all
+"Warning: All-all was deprecated in v0.16.
+Please use all⁻ instead."
+#-}
+all-All = all⁺
+{-# WARNING_ON_USAGE all-All
+"Warning: all-All was deprecated in v0.16.
+Please use all⁺ instead."
+#-}
+All-map = map⁺
+{-# WARNING_ON_USAGE All-map
+"Warning: All-map was deprecated in v0.16.
+Please use map⁺ instead."
+#-}
+map-All = map⁻
+{-# WARNING_ON_USAGE map-All
+"Warning: map-All was deprecated in v0.16.
+Please use map⁻ instead."
+#-}
+
diff --git a/src/Data/List/Any.agda b/src/Data/List/Any.agda
index a9f94da..1d3f106 100644
--- a/src/Data/List/Any.agda
+++ b/src/Data/List/Any.agda
@@ -8,47 +8,69 @@ module Data.List.Any {a} {A : Set a} where
open import Data.Empty
open import Data.Fin
-open import Data.List.Base as List using (List; []; _∷_)
+open import Data.List.Base as List using (List; []; [_]; _∷_)
open import Data.Product as Prod using (∃; _,_)
+open import Data.Sum as Sum using (_⊎_; inj₁; inj₂)
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.Nullary.Negation using (contradiction)
+open import Relation.Unary hiding (_∈_)
+------------------------------------------------------------------------
-- Any P xs means that at least one element in xs satisfies P.
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 : ∀ {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.
+------------------------------------------------------------------------
+-- Operations on Any
-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
+module _ {p} {P : A → Set p} {x xs} where
--- Decides Any.
+ head : ¬ Any P xs → Any P (x ∷ xs) → P x
+ head ¬pxs (here px) = px
+ head ¬pxs (there pxs) = contradiction pxs ¬pxs
-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)
-any p (x ∷ xs) | no ¬px = Dec.map′ there (tail ¬px) (any p xs)
+ tail : ¬ P x → Any P (x ∷ xs) → Any P xs
+ tail ¬px (here px) = ⊥-elim (¬px px)
+ tail ¬px (there pxs) = pxs
--- index x∈xs is the list position (zero-based) which x∈xs points to.
+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)
+-- `index x∈xs` is the list position (zero-based) which `x∈xs` points to.
index : ∀ {p} {P : A → Set p} {xs} → Any P xs → Fin (List.length xs)
index (here px) = zero
index (there pxs) = suc (index pxs)
-- If any element satisfies P, then P is satisfied.
-
satisfied : ∀ {p} {P : A → Set p} {xs} → Any P xs → ∃ P
satisfied (here px) = _ , px
satisfied (there pxs) = satisfied pxs
+
+module _ {p} {P : A → Set p} {x xs} where
+
+ toSum : Any P (x ∷ xs) → P x ⊎ Any P xs
+ toSum (here px) = inj₁ px
+ toSum (there pxs) = inj₂ pxs
+
+ fromSum : P x ⊎ Any P xs → Any P (x ∷ xs)
+ fromSum (inj₁ px) = here px
+ fromSum (inj₂ pxs) = there pxs
+
+------------------------------------------------------------------------
+-- Properties of predicates preserved by Any
+
+module _ {p} {P : A → Set p} where
+
+ any : Decidable P → Decidable (Any P)
+ any P? [] = no λ()
+ any P? (x ∷ xs) with P? x
+ ... | yes px = yes (here px)
+ ... | no ¬px = Dec.map′ there (tail ¬px) (any P? xs)
+
+ satisfiable : Satisfiable P → Satisfiable (Any P)
+ satisfiable (x , Px) = [ x ] , here Px
diff --git a/src/Data/List/Any/BagAndSetEquality.agda b/src/Data/List/Any/BagAndSetEquality.agda
deleted file mode 100644
index ad8bbba..0000000
--- a/src/Data/List/Any/BagAndSetEquality.agda
+++ /dev/null
@@ -1,274 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- Properties related to bag and set equality
-------------------------------------------------------------------------
-
--- Bag and set equality are defined in Data.List.Any.
-
-module Data.List.Any.BagAndSetEquality where
-
-open import Algebra
-open import Algebra.FunctionProperties
-open import Category.Monad
-open import Data.List as List
-import Data.List.Properties as LP
-open import Data.List.Any as Any using (Any); open 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
-open import Function.Equality using (_⟨$⟩_)
-import Function.Equivalence as FE
-open import Function.Inverse as Inv using (_↔_; module Inverse)
-open import Function.Related as Related using (↔⇒; ⌊_⌋; ⌊_⌋→; ⇒→)
-open import Function.Related.TypeIsomorphisms
-open import Relation.Binary
-import Relation.Binary.EqReasoning as EqR
-open import Relation.Binary.PropositionalEquality as P
- using (_≡_; _≗_)
-open import Relation.Binary.Sum
-open import Relation.Nullary
-
-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)
- module ×⊎ {k ℓ} = CommutativeSemiring (×⊎-CommutativeSemiring k ℓ)
- open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
- module ListMonoid {a} {A : Set a} = Monoid (List.monoid A)
-
-------------------------------------------------------------------------
--- Congruence lemmas
-
--- _∷_ is a congruence.
-
-∷-cong : ∀ {a k} {A : Set a} {x₁ x₂ : A} {xs₁ xs₂} →
- x₁ ≡ x₂ → xs₁ ∼[ k ] xs₂ → x₁ ∷ xs₁ ∼[ k ] x₂ ∷ xs₂
-∷-cong {x₂ = x} {xs₁} {xs₂} P.refl xs₁≈xs₂ {y} =
- y ∈ x ∷ xs₁ ↔⟨ sym $ ∷↔ (_≡_ y) ⟩
- (y ≡ x ⊎ y ∈ xs₁) ∼⟨ (y ≡ x ∎) ⊎-cong xs₁≈xs₂ ⟩
- (y ≡ x ⊎ y ∈ xs₂) ↔⟨ ∷↔ (_≡_ y) ⟩
- y ∈ x ∷ xs₂ ∎
- where open Related.EquationalReasoning
-
--- List.map is a congruence.
-
-map-cong : ∀ {ℓ k} {A B : Set ℓ} {f₁ f₂ : A → B} {xs₁ xs₂} →
- f₁ ≗ f₂ → xs₁ ∼[ k ] xs₂ →
- List.map f₁ xs₁ ∼[ k ] List.map f₂ xs₂
-map-cong {ℓ} {f₁ = f₁} {f₂} {xs₁} {xs₂} f₁≗f₂ xs₁≈xs₂ {x} =
- x ∈ List.map f₁ xs₁ ↔⟨ sym $ map↔ {a = ℓ} {b = ℓ} {p = ℓ} ⟩
- Any (λ y → x ≡ f₁ y) xs₁ ∼⟨ Any-cong (↔⇒ ∘ helper) xs₁≈xs₂ ⟩
- Any (λ y → x ≡ f₂ y) xs₂ ↔⟨ map↔ {a = ℓ} {b = ℓ} {p = ℓ} ⟩
- x ∈ List.map f₂ xs₂ ∎
- where
- open Related.EquationalReasoning
-
- helper : ∀ y → x ≡ f₁ y ↔ x ≡ f₂ y
- helper y = record
- { to = P.→-to-⟶ (λ x≡f₁y → P.trans x≡f₁y ( f₁≗f₂ y))
- ; from = P.→-to-⟶ (λ x≡f₂y → P.trans x≡f₂y (P.sym $ f₁≗f₂ y))
- ; inverse-of = record
- { left-inverse-of = λ _ → P.proof-irrelevance _ _
- ; right-inverse-of = λ _ → P.proof-irrelevance _ _
- }
- }
-
--- _++_ is a congruence.
-
-++-cong : ∀ {a k} {A : Set a} {xs₁ xs₂ ys₁ ys₂ : List A} →
- xs₁ ∼[ k ] xs₂ → ys₁ ∼[ k ] ys₂ →
- xs₁ ++ ys₁ ∼[ k ] xs₂ ++ ys₂
-++-cong {a} {xs₁ = xs₁} {xs₂} {ys₁} {ys₂} xs₁≈xs₂ ys₁≈ys₂ {x} =
- x ∈ xs₁ ++ ys₁ ↔⟨ sym $ ++↔ {a = a} {p = a} ⟩
- (x ∈ xs₁ ⊎ x ∈ ys₁) ∼⟨ xs₁≈xs₂ ⊎-cong ys₁≈ys₂ ⟩
- (x ∈ xs₂ ⊎ x ∈ ys₂) ↔⟨ ++↔ {a = a} {p = a} ⟩
- x ∈ xs₂ ++ ys₂ ∎
- where open Related.EquationalReasoning
-
--- concat is a congruence.
-
-concat-cong : ∀ {a k} {A : Set a} {xss₁ xss₂ : List (List A)} →
- xss₁ ∼[ k ] xss₂ → concat xss₁ ∼[ k ] concat xss₂
-concat-cong {a} {xss₁ = xss₁} {xss₂} xss₁≈xss₂ {x} =
- x ∈ concat xss₁ ↔⟨ sym $ concat↔ {a = a} {p = a} ⟩
- Any (Any (_≡_ x)) xss₁ ∼⟨ Any-cong (λ _ → _ ∎) xss₁≈xss₂ ⟩
- Any (Any (_≡_ x)) xss₂ ↔⟨ concat↔ {a = a} {p = a} ⟩
- x ∈ concat xss₂ ∎
- where open Related.EquationalReasoning
-
--- The list monad's bind is a congruence.
-
->>=-cong : ∀ {ℓ k} {A B : Set ℓ} {xs₁ xs₂} {f₁ f₂ : A → List B} →
- xs₁ ∼[ k ] xs₂ → (∀ x → f₁ x ∼[ k ] f₂ x) →
- (xs₁ >>= f₁) ∼[ k ] (xs₂ >>= f₂)
->>=-cong {ℓ} {xs₁ = xs₁} {xs₂} {f₁} {f₂} xs₁≈xs₂ f₁≈f₂ {x} =
- x ∈ (xs₁ >>= f₁) ↔⟨ sym $ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
- Any (λ y → x ∈ f₁ y) xs₁ ∼⟨ Any-cong (λ x → f₁≈f₂ x) xs₁≈xs₂ ⟩
- Any (λ y → x ∈ f₂ y) xs₂ ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
- x ∈ (xs₂ >>= f₂) ∎
- where open Related.EquationalReasoning
-
--- _⊛_ is a congruence.
-
-⊛-cong : ∀ {ℓ k} {A B : Set ℓ} {fs₁ fs₂ : List (A → B)} {xs₁ xs₂} →
- fs₁ ∼[ k ] fs₂ → xs₁ ∼[ k ] xs₂ →
- (fs₁ ⊛ xs₁) ∼[ k ] (fs₂ ⊛ xs₂)
-⊛-cong fs₁≈fs₂ xs₁≈xs₂ =
- >>=-cong fs₁≈fs₂ λ f →
- >>=-cong xs₁≈xs₂ λ x →
- _ ∎
- where open Related.EquationalReasoning
-
--- _⊗_ is a congruence.
-
-⊗-cong : ∀ {ℓ k} {A B : Set ℓ} {xs₁ xs₂ : List A} {ys₁ ys₂ : List B} →
- xs₁ ∼[ k ] xs₂ → ys₁ ∼[ k ] ys₂ →
- (xs₁ ⊗ ys₁) ∼[ k ] (xs₂ ⊗ ys₂)
-⊗-cong {ℓ} xs₁≈xs₂ ys₁≈ys₂ =
- ⊛-cong (⊛-cong (Ord.refl {x = [ _,_ {a = ℓ} {b = ℓ} ]})
- xs₁≈xs₂)
- ys₁≈ys₂
-
-------------------------------------------------------------------------
--- Other properties
-
--- _++_ and [] form a commutative monoid, with either bag or set
--- equality as the underlying equality.
-
-commutativeMonoid : ∀ {a} → Symmetric-kind → Set a →
- CommutativeMonoid _ _
-commutativeMonoid {a} k A = record
- { Carrier = List A
- ; _≈_ = λ xs ys → xs ∼[ ⌊ k ⌋ ] ys
- ; _∙_ = _++_
- ; ε = []
- ; isCommutativeMonoid = record
- { isSemigroup = record
- { isEquivalence = Eq.isEquivalence
- ; assoc = λ xs ys zs →
- Eq.reflexive (ListMonoid.assoc xs ys zs)
- ; ∙-cong = ++-cong
- }
- ; identityˡ = λ xs {x} → x ∈ xs ∎
- ; comm = λ xs ys {x} →
- x ∈ xs ++ ys ↔⟨ ++↔++ {a = a} {p = a} xs ys ⟩
- x ∈ ys ++ xs ∎
- }
- }
- where open Related.EquationalReasoning
-
--- The only list which is bag or set equal to the empty list (or a
--- subset or subbag of the list) is the empty list itself.
-
-empty-unique : ∀ {k a} {A : Set a} {xs : List A} →
- xs ∼[ ⌊ k ⌋→ ] [] → xs ≡ []
-empty-unique {xs = []} _ = P.refl
-empty-unique {xs = _ ∷ _} ∷∼[] with ⇒→ ∷∼[] (here P.refl)
-... | ()
-
--- _++_ is idempotent (under set equality).
-
-++-idempotent : ∀ {a} {A : Set a} →
- Idempotent (λ (xs ys : List A) → xs ∼[ set ] ys) _++_
-++-idempotent {a} xs {x} =
- x ∈ xs ++ xs ∼⟨ FE.equivalence ([ id , id ]′ ∘ _⟨$⟩_ (Inverse.from $ ++↔ {a = a} {p = a}))
- (_⟨$⟩_ (Inverse.to $ ++↔ {a = a} {p = a}) ∘ inj₁) ⟩
- x ∈ xs ∎
- where open Related.EquationalReasoning
-
--- The list monad's bind distributes from the left over _++_.
-
->>=-left-distributive :
- ∀ {ℓ} {A B : Set ℓ} (xs : List A) {f g : A → List B} →
- (xs >>= λ x → f x ++ g x) ∼[ bag ] (xs >>= f) ++ (xs >>= g)
->>=-left-distributive {ℓ} xs {f} {g} {y} =
- y ∈ (xs >>= λ x → f x ++ g x) ↔⟨ sym $ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
- Any (λ x → y ∈ f x ++ g x) xs ↔⟨ sym (Any-cong (λ _ → ++↔ {a = ℓ} {p = ℓ}) (_ ∎)) ⟩
- Any (λ x → y ∈ f x ⊎ y ∈ g x) xs ↔⟨ sym $ ⊎↔ {a = ℓ} {p = ℓ} {q = ℓ} ⟩
- (Any (λ x → y ∈ f x) xs ⊎ Any (λ x → y ∈ g x) xs) ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟨ ×⊎.+-cong {ℓ = ℓ} ⟩ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
- (y ∈ (xs >>= f) ⊎ y ∈ (xs >>= g)) ↔⟨ ++↔ {a = ℓ} {p = ℓ} ⟩
- y ∈ (xs >>= f) ++ (xs >>= g) ∎
- where open Related.EquationalReasoning
-
--- The same applies to _⊛_.
-
-⊛-left-distributive :
- ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) xs₁ xs₂ →
- (fs ⊛ (xs₁ ++ xs₂)) ∼[ bag ] (fs ⊛ xs₁) ++ (fs ⊛ xs₂)
-⊛-left-distributive {B = B} fs xs₁ xs₂ = begin
- fs ⊛ (xs₁ ++ xs₂) ≡⟨ P.refl ⟩
- (fs >>= λ f → xs₁ ++ xs₂ >>= return ∘ f) ≡⟨ (LP.Monad.cong (P.refl {x = fs}) λ f →
- LP.Monad.right-distributive xs₁ xs₂ (return ∘ f)) ⟩
- (fs >>= λ f → (xs₁ >>= return ∘ f) ++
- (xs₂ >>= return ∘ f)) ≈⟨ >>=-left-distributive fs ⟩
-
- (fs >>= λ f → xs₁ >>= return ∘ f) ++
- (fs >>= λ f → xs₂ >>= return ∘ f) ≡⟨ P.refl ⟩
-
- (fs ⊛ xs₁) ++ (fs ⊛ xs₂) ∎
- where open EqR ([ bag ]-Equality B)
-
-private
-
- -- If x ∷ xs is set equal to x ∷ ys, then xs and ys are not
- -- necessarily set equal.
-
- ¬-drop-cons :
- ∀ {a} {A : Set a} {x : A} →
- ¬ (∀ {xs ys} → x ∷ xs ∼[ set ] x ∷ ys → xs ∼[ set ] ys)
- ¬-drop-cons {x = x} drop-cons
- with FE.Equivalence.to x∼[] ⟨$⟩ here P.refl
- where
- x,x≈x : (x ∷ x ∷ []) ∼[ set ] [ x ]
- x,x≈x = ++-idempotent [ x ]
-
- x∼[] : [ x ] ∼[ set ] []
- x∼[] = drop-cons x,x≈x
- ... | ()
-
--- However, the corresponding property does hold for bag equality.
-
-drop-cons : ∀ {a} {A : Set a} {x : A} {xs ys} →
- x ∷ xs ∼[ bag ] x ∷ ys → xs ∼[ bag ] ys
-drop-cons {A = A} {x} {xs} {ys} x∷xs≈x∷ys {z} = record
- { to = P.→-to-⟶ $ f x∷xs≈x∷ys
- ; from = P.→-to-⟶ $ f $ Inv.sym x∷xs≈x∷ys
- ; inverse-of = record
- { left-inverse-of = f∘f x∷xs≈x∷ys
- ; right-inverse-of = f∘f $ Inv.sym x∷xs≈x∷ys
- }
- }
- where
- open Inverse
- open P.≡-Reasoning
-
- f : ∀ {xs ys z} → (z ∈ x ∷ xs) ↔ (z ∈ x ∷ ys) → z ∈ xs → z ∈ ys
- f inv z∈xs with to inv ⟨$⟩ there z∈xs | left-inverse-of inv (there z∈xs)
- f inv z∈xs | there z∈ys | left⁺ = z∈ys
- f inv z∈xs | here z≡x | left⁺ with to inv ⟨$⟩ here z≡x | left-inverse-of inv (here z≡x)
- f inv z∈xs | here z≡x | left⁺ | there z∈ys | left⁰ = z∈ys
- f inv z∈xs | here P.refl | left⁺ | here P.refl | left⁰ with begin
- here P.refl ≡⟨ P.sym left⁰ ⟩
- from inv ⟨$⟩ here P.refl ≡⟨ left⁺ ⟩
- there z∈xs ∎
- ... | ()
-
- f∘f : ∀ {xs ys z}
- (inv : (z ∈ x ∷ xs) ↔ (z ∈ x ∷ ys)) (p : z ∈ xs) →
- f (Inv.sym inv) (f inv p) ≡ p
- f∘f inv z∈xs with to inv ⟨$⟩ there z∈xs | left-inverse-of inv (there z∈xs)
- f∘f inv z∈xs | there z∈ys | left⁺ with from inv ⟨$⟩ there z∈ys | right-inverse-of inv (there z∈ys)
- f∘f inv z∈xs | there z∈ys | P.refl | .(there z∈xs) | _ = P.refl
- f∘f inv z∈xs | here z≡x | left⁺ with to inv ⟨$⟩ here z≡x | left-inverse-of inv (here z≡x)
- f∘f inv z∈xs | here z≡x | left⁺ | there z∈ys | left⁰ with from inv ⟨$⟩ there z∈ys | right-inverse-of inv (there z∈ys)
- f∘f inv z∈xs | here z≡x | left⁺ | there z∈ys | P.refl | .(here z≡x) | _ with from inv ⟨$⟩ here z≡x
- | right-inverse-of inv (here z≡x)
- f∘f inv z∈xs | here z≡x | P.refl | there z∈ys | P.refl | .(here z≡x) | _ | .(there z∈xs) | _ = P.refl
- f∘f inv z∈xs | here P.refl | left⁺ | here P.refl | left⁰ with begin
- here P.refl ≡⟨ P.sym left⁰ ⟩
- from inv ⟨$⟩ here P.refl ≡⟨ left⁺ ⟩
- there z∈xs ∎
- ... | ()
diff --git a/src/Data/List/Any/Membership.agda b/src/Data/List/Any/Membership.agda
deleted file mode 100644
index f46729b..0000000
--- a/src/Data/List/Any/Membership.agda
+++ /dev/null
@@ -1,55 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- List membership and some related definitions
-------------------------------------------------------------------------
-
-open import Relation.Binary hiding (Decidable)
-
-module Data.List.Any.Membership {c ℓ} (S : Setoid c ℓ) where
-
-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 Setoid S renaming (Carrier to A)
-
--- 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
-
--- 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 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 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
diff --git a/src/Data/List/Any/Membership/Properties.agda b/src/Data/List/Any/Membership/Properties.agda
deleted file mode 100644
index d49e1dc..0000000
--- a/src/Data/List/Any/Membership/Properties.agda
+++ /dev/null
@@ -1,72 +0,0 @@
-------------------------------------------------------------------------
--- 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
deleted file mode 100644
index 4d29544..0000000
--- a/src/Data/List/Any/Membership/Propositional.agda
+++ /dev/null
@@ -1,84 +0,0 @@
-------------------------------------------------------------------------
--- 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
deleted file mode 100644
index 6b86928..0000000
--- a/src/Data/List/Any/Membership/Propositional/Properties.agda
+++ /dev/null
@@ -1,286 +0,0 @@
-------------------------------------------------------------------------
--- 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 4cf1fa1..e16084e 100644
--- a/src/Data/List/Any/Properties.agda
+++ b/src/Data/List/Any/Properties.agda
@@ -9,53 +9,91 @@
module Data.List.Any.Properties where
-open import Algebra
open import Category.Monad
open import Data.Bool.Base using (Bool; false; true; T)
open import Data.Bool.Properties
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.Categorical using (monad)
open import Data.List.Any as Any using (Any; here; there)
-open import Data.List.Any.Membership.Propositional
+open import Data.List.Membership.Propositional
+open import Data.List.Membership.Propositional.Properties.Core
+ using (Any↔; find∘map; map∘find; lose∘find)
+open import Data.List.Relation.Pointwise
+ using (Pointwise; []; _∷_)
open import Data.Nat using (zero; suc; _<_; z≤n; s≤s)
+open import Data.Maybe as Maybe using (Maybe; just; nothing)
open import Data.Product as Prod
using (_×_; _,_; ∃; ∃₂; proj₁; proj₂; uncurry′)
+open import Data.Product.Properties
+open import Data.Product.Relation.Pointwise.NonDependent
+ using (_×-cong_)
+import Data.Product.Relation.Pointwise.Dependent as Σ
open import Data.Sum as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
+open import Data.Sum.Relation.Pointwise using (_⊎-cong_)
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.Equivalence using (_⇔_; equivalence; Equivalence)
+open import Function.Inverse as Inv using (_↔_; inverse; Inverse)
+open import Function.Related as Related using (Related; SK-sym)
open import Relation.Binary
-open import Relation.Binary.Product.Pointwise
open import Relation.Binary.PropositionalEquality as P
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 import Relation.Nullary using (¬_)
open Related.EquationalReasoning
+
private
- open module ListMonad {ℓ} = RawMonad (List.monad {ℓ = ℓ})
+ open module ListMonad {ℓ} = RawMonad (monad {ℓ = ℓ})
+
+------------------------------------------------------------------------
+-- Equality properties
+
+module _ {a p ℓ} {A : Set a} {P : A → Set p} {_≈_ : Rel A ℓ} where
+
+ lift-resp : P Respects _≈_ → (Any P) Respects (Pointwise _≈_)
+ 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)
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ here-injective : ∀ {x xs} {p q : P x} →
+ here {P = P} {xs = xs} p ≡ here q → p ≡ q
+ here-injective refl = refl
+
+ there-injective : ∀ {x xs} {p q : Any P xs} →
+ there {x = x} p ≡ there q → p ≡ q
+ there-injective refl = refl
+
+------------------------------------------------------------------------
+-- Misc
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ ¬Any[] : ¬ Any P []
+ ¬Any[] ()
------------------------------------------------------------------------
--- If a predicate P respects the underlying equality then Any P
--- respects the list equality.
+-- Any is a congruence
+
+module _ {a k p q} {A : Set a} {P : Pred A p} {Q : Pred A q} where
-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)
+ Any-cong : ∀ {xs ys : List A} →
+ (∀ x → Related k (P x) (Q x)) →
+ Preorder._∼_ (Related.InducedPreorder₂ k {A = A} _∈_) xs ys →
+ Related k (Any P xs) (Any Q ys)
+ Any-cong {xs} {ys} P↔Q xs≈ys =
+ Any P xs ↔⟨ SK-sym Any↔ ⟩
+ (∃ λ x → x ∈ xs × P x) ∼⟨ Σ.cong Inv.id (xs≈ys ×-cong P↔Q _) ⟩
+ (∃ λ x → x ∈ ys × Q x) ↔⟨ Any↔ ⟩
+ Any Q ys ∎
------------------------------------------------------------------------
--- Some lemmas related to map, find and lose
--- Any.map is functorial.
+-- map
map-id : ∀ {a p} {A : Set a} {P : A → Set p} (f : P ⋐ P) {xs} →
(∀ {x} (p : P x) → f p ≡ p) →
@@ -71,75 +109,6 @@ map-∘ : ∀ {a p q r}
map-∘ f g (here p) = refl
map-∘ f g (there p) = P.cong there $ map-∘ f g p
--- Lemmas relating map and find.
-
-map∘find : ∀ {a p} {A : Set a} {P : A → Set p} {xs}
- (p : Any P xs) → let p′ = find p in
- {f : _≡_ (proj₁ p′) ⋐ P} →
- f refl ≡ proj₂ (proj₂ p′) →
- Any.map f (proj₁ (proj₂ p′)) ≡ p
-map∘find (here p) hyp = P.cong here hyp
-map∘find (there p) hyp = P.cong there (map∘find p hyp)
-
-find∘map : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q}
- {xs : List A} (p : Any P xs) (f : P ⋐ Q) →
- find (Any.map f p) ≡ Prod.map id (Prod.map id f) (find p)
-find∘map (here p) f = refl
-find∘map (there p) f rewrite find∘map p f = refl
-
--- find satisfies a simple equality when the predicate is a
--- propositional equality.
-
-find-∈ : ∀ {a} {A : Set a} {x : A} {xs : List A} (x∈xs : x ∈ xs) →
- find x∈xs ≡ (x , x∈xs , refl)
-find-∈ (here refl) = refl
-find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl
-
--- 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
-
-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-⟶ ∃∈-Any
- ; from = P.→-to-⟶ (find {P = P})
- ; inverse-of = record
- { left-inverse-of = λ p →
- find∘lose P (proj₁ (proj₂ p)) (proj₂ (proj₂ p))
- ; right-inverse-of = lose∘find
- }
- }
-
-------------------------------------------------------------------------
--- Any is a congruence
-
-Any-cong : ∀ {k ℓ} {A : Set ℓ} {P₁ P₂ : A → Set ℓ} {xs₁ xs₂ : List A} →
- (∀ x → Related k (P₁ x) (P₂ x)) → xs₁ ∼[ k ] xs₂ →
- Related k (Any P₁ xs₁) (Any P₂ xs₂)
-Any-cong {P₁ = P₁} {P₂} {xs₁} {xs₂} P₁↔P₂ xs₁≈xs₂ =
- Any P₁ xs₁ ↔⟨ sym $ Any↔ {P = P₁} ⟩
- (∃ λ x → x ∈ xs₁ × P₁ x) ∼⟨ Σ.cong Inv.id (xs₁≈xs₂ ×-cong P₁↔P₂ _) ⟩
- (∃ λ x → x ∈ xs₂ × P₂ x) ↔⟨ Any↔ {P = P₂} ⟩
- Any P₂ xs₂ ∎
-
------------------------------------------------------------------------
-- Swapping
@@ -167,41 +136,20 @@ swap-invol (there 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} = record
- { to = P.→-to-⟶ swap
- ; from = P.→-to-⟶ swap
- ; inverse-of = record
- { left-inverse-of = swap-invol
- ; right-inverse-of = swap-invol
- }
- }
+swap↔ {P = P} = inverse swap swap swap-invol swap-invol
------------------------------------------------------------------------
-- Lemmas relating Any to ⊥
⊥↔Any⊥ : ∀ {a} {A : Set a} {xs : List A} → ⊥ ↔ Any (const ⊥) xs
-⊥↔Any⊥ {A = A} = record
- { to = P.→-to-⟶ (λ ())
- ; from = P.→-to-⟶ (λ p → from p)
- ; inverse-of = record
- { left-inverse-of = λ ()
- ; right-inverse-of = λ p → from p
- }
- }
+⊥↔Any⊥ {A = A} = inverse (λ()) (λ p → from p) (λ()) (λ p → from p)
where
from : {xs : List A} → Any (const ⊥) xs → ∀ {b} {B : Set b} → B
from (here ())
from (there p) = from p
-⊥↔Any[] : ∀ {a} {A : Set a} {P : A → Set} → ⊥ ↔ Any P []
-⊥↔Any[] = record
- { to = P.→-to-⟶ (λ ())
- ; from = P.→-to-⟶ (λ ())
- ; inverse-of = record
- { left-inverse-of = λ ()
- ; right-inverse-of = λ ()
- }
- }
+⊥↔Any[] : ∀ {a p} {A : Set a} {P : A → Set p} → ⊥ ↔ Any P []
+⊥↔Any[] = inverse (λ()) (λ()) (λ()) (λ())
------------------------------------------------------------------------
-- Lemmas relating Any to ⊤
@@ -223,7 +171,7 @@ module _ {a} {A : Set a} where
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⁻ _ _)
+ any⇔ = equivalence (any⁺ _) (any⁻ _ _)
------------------------------------------------------------------------
-- Sums commute with Any
@@ -239,29 +187,21 @@ module _ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} where
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
- }
- }
+ ⊎↔ = inverse Any-⊎⁺ Any-⊎⁻ from∘to to∘from
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
+ from∘to (inj₁ (here p)) = refl
+ from∘to (inj₁ (there p)) rewrite from∘to (inj₁ p) = refl
+ from∘to (inj₂ (here q)) = refl
+ from∘to (inj₂ (there q)) rewrite from∘to (inj₂ q) = 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 (here (inj₁ p)) = refl
+ to∘from (here (inj₂ q)) = 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
+ to∘from (there .(Any.map inj₁ p)) | inj₁ p | refl = refl
+ to∘from (there .(Any.map inj₂ q)) | inj₂ q | refl = refl
------------------------------------------------------------------------
-- Products "commute" with Any.
@@ -280,24 +220,15 @@ module _ {a b p q} {A : Set a} {B : Set b}
×↔ : ∀ {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
- }
- }
+ ×↔ {xs} {ys} = inverse Any-×⁺ Any-×⁻ from∘to to∘from
where
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
+ from∘to (p , q) rewrite
+ find∘map p (λ p → Any.map (λ q → (p , q)) q)
+ | find∘map q (λ q → proj₂ (proj₂ (find p)) , q)
+ | lose∘find p
+ | lose∘find q
+ = refl
to∘from : ∀ pq → Any-×⁺ (Any-×⁻ pq) ≡ pq
to∘from pq
@@ -314,8 +245,7 @@ module _ {a b p q} {A : Set a} {B : Set b}
= 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)
+ helper rewrite P.sym $ map-∘ (λ q → p , q)
(λ y → P.subst Q y q)
y∈ys
= lem₂ _ refl
@@ -323,43 +253,64 @@ module _ {a b p q} {A : Set a} {B : Set b}
------------------------------------------------------------------------
-- Invertible introduction (⁺) and elimination (⁻) rules for various
-- list functions
+
+------------------------------------------------------------------------
+-- map
+
+module _ {a p} {A : Set a} {P : Pred A p} where
+
+ singleton⁺ : ∀ {x} → P x → Any P [ x ]
+ singleton⁺ Px = here Px
+
+ singleton⁻ : ∀ {x} → Any P [ x ] → P x
+ singleton⁻ (here Px) = Px
+ singleton⁻ (there ())
+
------------------------------------------------------------------------
-- map
-module _ {a b} {A : Set a} {B : Set b} where
+module _ {a b} {A : Set a} {B : Set b} {f : A → B} where
- map⁺ : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
+ map⁺ : ∀ {p} {P : B → Set p} {xs} →
Any (P ∘ f) xs → Any P (List.map f xs)
map⁺ (here p) = here p
map⁺ (there p) = there $ map⁺ p
- map⁻ : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
+ map⁻ : ∀ {p} {P : B → Set p} {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⁻ : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
+ map⁺∘map⁻ : ∀ {p} {P : B → Set p} {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⁺ : ∀ {p} (P : B → Set p) {f : A → B} {xs} →
+ map⁻∘map⁺ : ∀ {p} (P : B → Set p) {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↔ : ∀ {p} {P : B → Set p} {f : A → B} {xs} →
+ map↔ : ∀ {p} {P : B → Set p} {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↔ = inverse map⁺ map⁻ (map⁻∘map⁺ _) map⁺∘map⁻
+
+------------------------------------------------------------------------
+-- mapMaybe
+
+module _ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
+ (f : A → Maybe B) where
+
+ mapMaybe⁺ : ∀ xs → Any (Maybe.Any P) (map f xs) →
+ Any P (mapMaybe f xs)
+ mapMaybe⁺ [] ()
+ mapMaybe⁺ (x ∷ xs) ps with f x | ps
+ ... | nothing | here ()
+ ... | nothing | there pxs = mapMaybe⁺ xs pxs
+ ... | just _ | here (just py) = here py
+ ... | just _ | there pxs = there (mapMaybe⁺ xs pxs)
------------------------------------------------------------------------
-- _++_
@@ -396,14 +347,7 @@ module _ {a p} {A : Set a} {P : A → Set p} where
++⁻∘++⁺ (x ∷ xs) (inj₂ p) rewrite ++⁻∘++⁺ xs (inj₂ p) = refl
++↔ : ∀ {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
- }
- }
+ ++↔ {xs = xs} = inverse [ ++⁺ˡ , ++⁺ʳ xs ]′ (++⁻ xs) (++⁻∘++⁺ xs) (++⁺∘++⁻ xs)
++-comm : ∀ xs ys → Any P (xs ++ ys) → Any P (ys ++ xs)
++-comm xs ys = [ ++⁺ʳ ys , ++⁺ˡ ]′ ∘ ++⁻ xs
@@ -411,28 +355,25 @@ module _ {a p} {A : Set a} {P : A → Set p} where
++-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
+ rewrite ++⁻∘++⁺ ys {ys = []} (inj₁ p) = refl
++-comm∘++-comm (x ∷ xs) {ys} (here p)
- rewrite ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₂ (here p)) = P.refl
+ rewrite ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₂ (here 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
+ | inj₁ p | refl
rewrite ++⁻∘++⁺ ys (inj₂ p)
- | ++⁻∘++⁺ ys (inj₂ $ there {x = x} p) = P.refl
+ | ++⁻∘++⁺ ys (inj₂ $ there {x = x} p) = refl
++-comm∘++-comm (x ∷ xs) {ys} (there .([ ++⁺ʳ xs , ++⁺ˡ ]′ (++⁻ ys (++⁺ˡ p))))
- | inj₂ p | P.refl
+ | inj₂ p | refl
rewrite ++⁻∘++⁺ ys {ys = xs} (inj₁ p)
- | ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₁ p) = P.refl
+ | ++⁻∘++⁺ ys {ys = x ∷ xs} (inj₁ 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
- }
- }
+ ++↔++ xs ys = inverse (++-comm xs ys) (++-comm ys xs)
+ (++-comm∘++-comm xs) (++-comm∘++-comm ys)
+
+ ++-insert : ∀ xs {ys x} → P x → Any P (xs ++ [ x ] ++ ys)
+ ++-insert xs Px = ++⁺ʳ xs (++⁺ˡ (singleton⁺ Px))
------------------------------------------------------------------------
-- concat
@@ -478,14 +419,7 @@ module _ {a p} {A : Set a} {P : A → Set p} where
P.cong there $ concat⁻∘concat⁺ p
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
- }
- }
+ concat↔ {xss} = inverse concat⁺ (concat⁻ xss) concat⁻∘concat⁺ (concat⁺∘concat⁻ xss)
------------------------------------------------------------------------
-- applyUpTo
@@ -541,16 +475,8 @@ 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-∈↔ = record
- { to = P.→-to-⟶ (map-with-∈⁺ _)
- ; from = P.→-to-⟶ (map-with-∈⁻ _ _)
- ; inverse-of = record
- { left-inverse-of = from∘to _
- ; right-inverse-of = to∘from _ _
- }
- }
+ map-with-∈↔ = inverse (map-with-∈⁺ _) (map-with-∈⁻ _ _) (from∘to _) (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
@@ -566,6 +492,7 @@ module _ {a b p} {A : Set a} {B : Set b} {P : B → Set p} where
to∘from (y ∷ xs) f (there p) =
P.cong there $ to∘from xs (f ∘ there) p
+
------------------------------------------------------------------------
-- return
@@ -587,69 +514,67 @@ module _ {a p} {A : Set a} {P : A → Set p} where
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⁻
- }
- }
+ return↔ = inverse return⁺ return⁻ return⁻∘return⁺ return⁺∘return⁻
+
+------------------------------------------------------------------------
+-- _∷_
+
+module _ {a p} {A : Set a} where
+
+ ∷↔ : ∀ (P : Pred A 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) ∎
+
+------------------------------------------------------------------------
+-- _>>=_
+
+module _ {ℓ p} {A B : Set ℓ} {P : B → Set p} {f : A → List B} where
+
+ >>=↔ : ∀ {xs} → Any (Any P ∘ f) xs ↔ Any P (xs >>= f)
+ >>=↔ {xs} =
+ Any (Any P ∘ f) xs ↔⟨ map↔ ⟩
+ Any (Any P) (List.map f xs) ↔⟨ concat↔ ⟩
+ Any P (xs >>= f) ∎
+
+------------------------------------------------------------------------
+-- _⊛_
+
+module _ {ℓ} {A B : Set ℓ} where
+
+ ⊛↔ : ∀ {P : B → Set ℓ} {fs : List (A → B)} {xs : List A} →
+ Any (λ f → Any (P ∘ f) xs) fs ↔ Any P (fs ⊛ xs)
+ ⊛↔ {P = P} {fs} {xs} =
+ Any (λ f → Any (P ∘ f) xs) fs ↔⟨ Any-cong (λ _ → Any-cong (λ _ → return↔) (_ ∎)) (_ ∎) ⟩
+ Any (λ f → Any (Any P ∘ return ∘ f) xs) fs ↔⟨ Any-cong (λ _ → >>=↔ ) (_ ∎) ⟩
+ Any (λ f → Any P (xs >>= return ∘ f)) fs ↔⟨ >>=↔ ⟩
+ Any P (fs ⊛ xs) ∎
+
+-- An alternative introduction rule for _⊛_
+
+ ⊛⁺′ : ∀ {P : A → Set ℓ} {Q : B → Set ℓ} {fs : List (A → B)} {xs} →
+ Any (P ⟨→⟩ Q) fs → Any P xs → Any Q (fs ⊛ xs)
+ ⊛⁺′ pq p =
+ Inverse.to ⊛↔ ⟨$⟩
+ Any.map (λ pq → Any.map (λ {x} → pq {x}) p) pq
------------------------------------------------------------------------
--- _∷_.
-
-∷↔ : ∀ {a 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) ∎
-
--- _>>=_.
-
->>=↔ : ∀ {ℓ p} {A B : Set ℓ} {P : B → Set p} {xs} {f : A → List B} →
- Any (Any P ∘ f) xs ↔ Any P (xs >>= f)
->>=↔ {P = P} {xs} {f} =
- Any (Any P ∘ f) xs ↔⟨ map↔ {P = Any P} {f = f} ⟩
- Any (Any P) (List.map f xs) ↔⟨ concat↔ {P = P} ⟩
- Any P (xs >>= f) ∎
-
--- _⊛_.
-
-⊛↔ : ∀ {ℓ} {A B : Set ℓ} {P : B → Set ℓ}
- {fs : List (A → B)} {xs : List A} →
- Any (λ f → Any (P ∘ f) xs) fs ↔ Any P (fs ⊛ xs)
-⊛↔ {ℓ} {P = P} {fs} {xs} =
- Any (λ f → Any (P ∘ f) xs) fs ↔⟨ Any-cong (λ _ → Any-cong (λ _ → return↔ {a = ℓ} {p = ℓ}) (_ ∎)) (_ ∎) ⟩
- Any (λ f → Any (Any P ∘ return ∘ f) xs) fs ↔⟨ Any-cong (λ _ → >>=↔ {ℓ = ℓ} {p = ℓ}) (_ ∎) ⟩
- Any (λ f → Any P (xs >>= return ∘ f)) fs ↔⟨ >>=↔ {ℓ = ℓ} {p = ℓ} ⟩
- Any P (fs ⊛ xs) ∎
-
--- An alternative introduction rule for _⊛_.
-
-⊛⁺′ : ∀ {ℓ} {A B : Set ℓ} {P : A → Set ℓ} {Q : B → Set ℓ}
- {fs : List (A → B)} {xs} →
- Any (P ⟨→⟩ Q) fs → Any P xs → Any Q (fs ⊛ xs)
-⊛⁺′ {ℓ} pq p =
- Inverse.to (⊛↔ {ℓ = ℓ}) ⟨$⟩
- Any.map (λ pq → Any.map (λ {x} → pq {x}) p) pq
-
--- _⊗_.
-
-⊗↔ : ∀ {ℓ} {A B : Set ℓ} {P : A × B → Set ℓ}
- {xs : List A} {ys : List B} →
- Any (λ x → Any (λ y → P (x , y)) ys) xs ↔ Any P (xs ⊗ ys)
-⊗↔ {ℓ} {P = P} {xs} {ys} =
- Any (λ x → Any (λ y → P (x , y)) ys) xs ↔⟨ return↔ {a = ℓ} {p = ℓ} ⟩
- Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (return _,_) ↔⟨ ⊛↔ ⟩
- Any (λ x, → Any (P ∘ x,) ys) (_,_ <$> xs) ↔⟨ ⊛↔ ⟩
- Any P (xs ⊗ ys) ∎
-
-⊗↔′ : {A B : Set} {P : A → Set} {Q : B → Set}
- {xs : List A} {ys : List B} →
- (Any P xs × Any Q ys) ↔ Any (P ⟨×⟩ Q) (xs ⊗ ys)
-⊗↔′ {P = P} {Q} {xs} {ys} =
- (Any P xs × Any Q ys) ↔⟨ ×↔ ⟩
- Any (λ x → Any (λ y → P x × Q y) ys) xs ↔⟨ ⊗↔ ⟩
- Any (P ⟨×⟩ Q) (xs ⊗ ys) ∎
+-- _⊗_
+
+module _ {ℓ} {A B : Set ℓ} where
+
+ ⊗↔ : {P : A × B → Set ℓ} {xs : List A} {ys : List B} →
+ Any (λ x → Any (λ y → P (x , y)) ys) xs ↔ Any P (xs ⊗ ys)
+ ⊗↔ {P} {xs} {ys} =
+ Any (λ x → Any (λ y → P (x , y)) ys) xs ↔⟨ return↔ ⟩
+ Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (return _,_) ↔⟨ ⊛↔ ⟩
+ Any (λ x, → Any (P ∘ x,) ys) (_,_ <$> xs) ↔⟨ ⊛↔ ⟩
+ Any P (xs ⊗ ys) ∎
+
+ ⊗↔′ : {P : A → Set ℓ} {Q : B → Set ℓ} {xs : List A} {ys : List B} →
+ (Any P xs × Any Q ys) ↔ Any (P ⟨×⟩ Q) (xs ⊗ ys)
+ ⊗↔′ {P} {Q} {xs} {ys} =
+ (Any P xs × Any Q ys) ↔⟨ ×↔ ⟩
+ Any (λ x → Any (λ y → P x × Q y) ys) xs ↔⟨ ⊗↔ ⟩
+ Any (P ⟨×⟩ Q) (xs ⊗ ys) ∎
diff --git a/src/Data/List/Base.agda b/src/Data/List/Base.agda
index 168a93a..76fa75c 100644
--- a/src/Data/List/Base.agda
+++ b/src/Data/List/Base.agda
@@ -9,10 +9,15 @@ 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.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 using (id; _∘_)
+open import Data.These using (These; this; that; these)
+open import Function using (id; _∘_ ; _∘′_)
+open import Relation.Nullary using (yes; no)
+open import Relation.Unary using (Pred; Decidable)
+open import Relation.Unary.Properties using (∁?)
------------------------------------------------------------------------
-- Types
@@ -21,54 +26,70 @@ open import Agda.Builtin.List public
using (List; []; _∷_)
------------------------------------------------------------------------
--- Some operations
+-- Operations for transforming lists
--- * Basic functions
+map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → List A → List B
+map f [] = []
+map f (x ∷ xs) = f x ∷ map f xs
-infixr 5 _++_
+mapMaybe : ∀ {a b} {A : Set a} {B : Set b} → (A → Maybe B) → List A → List B
+mapMaybe p [] = []
+mapMaybe p (x ∷ xs) with p x
+... | just y = y ∷ mapMaybe p xs
+... | nothing = mapMaybe p xs
-[_] : ∀ {a} {A : Set a} → A → List A
-[ x ] = x ∷ []
+infixr 5 _++_
_++_ : ∀ {a} {A : Set a} → List A → List A → List A
[] ++ ys = ys
(x ∷ xs) ++ ys = x ∷ (xs ++ ys)
--- Snoc.
+intersperse : ∀ {a} {A : Set a} → A → List A → List A
+intersperse x [] = []
+intersperse x (y ∷ []) = y ∷ []
+intersperse x (y ∷ ys) = y ∷ x ∷ intersperse x ys
-infixl 5 _∷ʳ_
+------------------------------------------------------------------------
+-- Aligning and Zipping
-_∷ʳ_ : ∀ {a} {A : Set a} → List A → A → List A
-xs ∷ʳ x = xs ++ [ x ]
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
-null : ∀ {a} {A : Set a} → List A → Bool
-null [] = true
-null (x ∷ xs) = false
+ alignWith : (These A B → C) → List A → List B → List C
+ alignWith f [] bs = map (f ∘′ that) bs
+ alignWith f as [] = map (f ∘′ this) as
+ alignWith f (a ∷ as) (b ∷ bs) = f (these a b) ∷ alignWith f as bs
--- * List transformations
+ zipWith : (A → B → C) → List A → List B → List C
+ zipWith f (x ∷ xs) (y ∷ ys) = f x y ∷ zipWith f xs ys
+ zipWith f _ _ = []
-map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → List A → List B
-map f [] = []
-map f (x ∷ xs) = f x ∷ map f xs
+ unalignWith : (A → These B C) → List A → List B × List C
+ unalignWith f [] = [] , []
+ unalignWith f (a ∷ as) with f a
+ ... | this b = Prod.map₁ (b ∷_) (unalignWith f as)
+ ... | that c = Prod.map₂ (c ∷_) (unalignWith f as)
+ ... | these b c = Prod.map (b ∷_) (c ∷_) (unalignWith f as)
-replicate : ∀ {a} {A : Set a} → (n : ℕ) → A → List A
-replicate zero x = []
-replicate (suc n) x = x ∷ replicate n x
+ unzipWith : (A → B × C) → List A → List B × List C
+ unzipWith f [] = [] , []
+ unzipWith f (xy ∷ xys) = Prod.zip _∷_ _∷_ (f xy) (unzipWith f xys)
-zipWith : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
- → (A → B → C) → List A → List B → List C
-zipWith f (x ∷ xs) (y ∷ ys) = f x y ∷ zipWith f xs ys
-zipWith f _ _ = []
+module _ {a b} {A : Set a} {B : Set b} where
-zip : ∀ {a b} {A : Set a} {B : Set b} → List A → List B → List (A × B)
-zip = zipWith (_,_)
+ align : List A → List B → List (These A B)
+ align = alignWith id
-intersperse : ∀ {a} {A : Set a} → A → List A → List A
-intersperse x [] = []
-intersperse x (y ∷ []) = [ y ]
-intersperse x (y ∷ z ∷ zs) = y ∷ x ∷ intersperse x (z ∷ zs)
+ zip : List A → List B → List (A × B)
+ zip = zipWith (_,_)
+
+ unalign : List (These A B) → List A × List B
+ unalign = unalignWith id
--- * Reducing lists (folds)
+ unzip : List (A × B) → List A × List B
+ unzip = unzipWith id
+
+------------------------------------------------------------------------
+-- Operations for reducing lists
foldr : ∀ {a b} {A : Set a} {B : Set b} → (A → B → B) → B → List A → B
foldr c n [] = n
@@ -78,8 +99,6 @@ foldl : ∀ {a b} {A : Set a} {B : Set b} → (A → B → A) → A → List B
foldl c n [] = n
foldl c n (x ∷ xs) = foldl c (c n x) xs
--- ** Special folds
-
concat : ∀ {a} {A : Set a} → List (List A) → List A
concat = foldr _++_ []
@@ -87,6 +106,10 @@ concatMap : ∀ {a b} {A : Set a} {B : Set b} →
(A → List B) → List A → List B
concatMap f = concat ∘ map f
+null : ∀ {a} {A : Set a} → List A → Bool
+null [] = true
+null (x ∷ xs) = false
+
and : List Bool → Bool
and = foldr _∧_ true
@@ -108,12 +131,29 @@ product = foldr _*_ 1
length : ∀ {a} {A : Set a} → List A → ℕ
length = foldr (λ _ → suc) 0
-reverse : ∀ {a} {A : Set a} → List A → List A
-reverse = foldl (λ rev x → x ∷ rev) []
+------------------------------------------------------------------------
+-- Operations for constructing lists
+
+[_] : ∀ {a} {A : Set a} → A → List A
+[ x ] = x ∷ []
+
+fromMaybe : ∀ {a} {A : Set a} → Maybe A → List A
+fromMaybe (just x) = [ x ]
+fromMaybe nothing = []
+
+replicate : ∀ {a} {A : Set a} → (n : ℕ) → A → List A
+replicate zero x = []
+replicate (suc n) x = x ∷ replicate n x
+
+inits : ∀ {a} {A : Set a} → List A → List (List A)
+inits [] = [] ∷ []
+inits (x ∷ xs) = [] ∷ map (x ∷_) (inits xs)
--- * Building lists
+tails : ∀ {a} {A : Set a} → List A → List (List A)
+tails [] = [] ∷ []
+tails (x ∷ xs) = (x ∷ xs) ∷ tails xs
--- ** Scans
+-- Scans
scanr : ∀ {a b} {A : Set a} {B : Set b} →
(A → B → B) → B → List A → List B
@@ -127,58 +167,66 @@ scanl : ∀ {a b} {A : Set a} {B : Set b} →
scanl f e [] = e ∷ []
scanl f e (x ∷ xs) = e ∷ scanl f (f e x) xs
--- ** Unfolding
-
--- Unfold. Uses a measure (a natural number) to ensure termination.
-
-unfold : ∀ {a b} {A : Set a} (B : ℕ → Set b)
- (f : ∀ {n} → B (suc n) → Maybe (A × B n)) →
- ∀ {n} → B n → List A
-unfold B f {n = zero} s = []
-unfold B f {n = suc n} s with f s
-... | nothing = []
-... | just (x , s') = x ∷ unfold B f s'
-
--- applyUpTo 3 = f0 ∷ f1 ∷ f2 ∷ [].
+-- Tabulation
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 = 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)
+lookup : ∀ {a} {A : Set a} (xs : List A) → Fin (length xs) → A
+lookup [] ()
+lookup (x ∷ xs) fzero = x
+lookup (x ∷ xs) (fsuc i) = lookup xs i
+
+-- Numerical
+
+upTo : ℕ → List ℕ
+upTo = applyUpTo id
+
+downFrom : ℕ → List ℕ
+downFrom = applyDownFrom id
+
allFin : ∀ n → List (Fin n)
allFin n = tabulate id
--- ** Conversions
+-- Other
-fromMaybe : ∀ {a} {A : Set a} → Maybe A → List A
-fromMaybe (just x) = [ x ]
-fromMaybe nothing = []
+unfold : ∀ {a b} {A : Set a} (B : ℕ → Set b)
+ (f : ∀ {n} → B (suc n) → Maybe (A × B n)) →
+ ∀ {n} → B n → List A
+unfold B f {n = zero} s = []
+unfold B f {n = suc n} s with f s
+... | nothing = []
+... | just (x , s') = x ∷ unfold B f s'
+
+------------------------------------------------------------------------
+-- Operations for deconstructing lists
+
+-- Note that although these combinators can be useful for programming, when
+-- proving it is often a better idea to manually destruct a list argument:
+-- each branch of the pattern-matching will have a refined type.
--- * Sublists
+module _ {a} {A : Set a} where
--- ** Extracting sublists
+ uncons : List A → Maybe (A × List A)
+ uncons [] = nothing
+ uncons (x ∷ xs) = just (x , xs)
+
+ head : List A → Maybe A
+ head [] = nothing
+ head (x ∷ _) = just x
+
+ tail : List A → Maybe (List A)
+ tail [] = nothing
+ tail (_ ∷ xs) = just xs
take : ∀ {a} {A : Set a} → ℕ → List A → List A
take zero xs = []
@@ -196,34 +244,59 @@ splitAt (suc n) [] = ([] , [])
splitAt (suc n) (x ∷ xs) with splitAt n xs
... | (ys , zs) = (x ∷ ys , zs)
-takeWhile : ∀ {a} {A : Set a} → (A → Bool) → List A → List A
-takeWhile p [] = []
-takeWhile p (x ∷ xs) with p x
-... | true = x ∷ takeWhile p xs
-... | false = []
+takeWhile : ∀ {a p} {A : Set a} {P : Pred A p} →
+ Decidable P → List A → List A
+takeWhile P? [] = []
+takeWhile P? (x ∷ xs) with P? x
+... | yes _ = x ∷ takeWhile P? xs
+... | no _ = []
+
+dropWhile : ∀ {a p} {A : Set a} {P : Pred A p} →
+ Decidable P → List A → List A
+dropWhile P? [] = []
+dropWhile P? (x ∷ xs) with P? x
+... | yes _ = dropWhile P? xs
+... | no _ = x ∷ xs
+
+filter : ∀ {a p} {A : Set a} {P : Pred A p} →
+ Decidable P → List A → List A
+filter P? [] = []
+filter P? (x ∷ xs) with P? x
+... | no _ = filter P? xs
+... | yes _ = x ∷ filter P? xs
+
+partition : ∀ {a p} {A : Set a} {P : Pred A p} →
+ Decidable P → List A → (List A × List A)
+partition P? [] = ([] , [])
+partition P? (x ∷ xs) with P? x | partition P? xs
+... | yes _ | (ys , zs) = (x ∷ ys , zs)
+... | no _ | (ys , zs) = (ys , x ∷ zs)
+
+span : ∀ {a p} {A : Set a} {P : Pred A p} →
+ Decidable P → List A → (List A × List A)
+span P? [] = ([] , [])
+span P? (x ∷ xs) with P? x
+... | yes _ = Prod.map (x ∷_) id (span P? xs)
+... | no _ = ([] , x ∷ xs)
+
+break : ∀ {a p} {A : Set a} {P : Pred A p} →
+ Decidable P → List A → (List A × List A)
+break P? = span (∁? P?)
-dropWhile : ∀ {a} {A : Set a} → (A → Bool) → List A → List A
-dropWhile p [] = []
-dropWhile p (x ∷ xs) with p x
-... | true = dropWhile p xs
-... | false = x ∷ xs
+------------------------------------------------------------------------
+-- Operations for reversing lists
-span : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A)
-span p [] = ([] , [])
-span p (x ∷ xs) with p x
-... | true = Prod.map (_∷_ x) id (span p xs)
-... | false = ([] , x ∷ xs)
+reverse : ∀ {a} {A : Set a} → List A → List A
+reverse = foldl (λ rev x → x ∷ rev) []
-break : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A)
-break p = span (not ∘ p)
+-- Snoc.
-inits : ∀ {a} {A : Set a} → List A → List (List A)
-inits [] = [] ∷ []
-inits (x ∷ xs) = [] ∷ map (_∷_ x) (inits xs)
+infixl 5 _∷ʳ_
-tails : ∀ {a} {A : Set a} → List A → List (List A)
-tails [] = [] ∷ []
-tails (x ∷ xs) = (x ∷ xs) ∷ tails xs
+_∷ʳ_ : ∀ {a} {A : Set a} → List A → A → List A
+xs ∷ʳ x = xs ++ [ x ]
+
+-- Backwards initialisation
infixl 5 _∷ʳ'_
@@ -231,30 +304,57 @@ data InitLast {a} {A : Set a} : List A → Set a where
[] : InitLast []
_∷ʳ'_ : (xs : List A) (x : A) → InitLast (xs ∷ʳ x)
-initLast : ∀ {a} {A : Set a} (xs : List A) → InitLast xs
+initLast : ∀ {a} {A : Set a} → (xs : List A) → InitLast xs
initLast [] = []
initLast (x ∷ xs) with initLast xs
initLast (x ∷ .[]) | [] = [] ∷ʳ' x
initLast (x ∷ .(ys ∷ʳ y)) | ys ∷ʳ' y = (x ∷ ys) ∷ʳ' y
--- * Searching lists
-
--- ** Searching with a predicate
+------------------------------------------------------------------------
+-- DEPRECATED
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+--
+-- Note that the `boolX` functions are not given warnings as they are
+-- used by other deprecated proofs throughout the library.
--- A generalised variant of filter.
+-- Version 0.15
-gfilter : ∀ {a b} {A : Set a} {B : Set b} →
- (A → Maybe B) → List A → List B
-gfilter p [] = []
-gfilter p (x ∷ xs) with p x
-... | just y = y ∷ gfilter p xs
-... | nothing = gfilter p xs
+gfilter = mapMaybe
+{-# WARNING_ON_USAGE gfilter
+"Warning: gfilter was deprecated in v0.15.
+Please use mapMaybe instead."
+#-}
-filter : ∀ {a} {A : Set a} → (A → Bool) → List A → List A
-filter p = gfilter (λ x → if p x then just x else nothing)
+boolFilter : ∀ {a} {A : Set a} → (A → Bool) → List A → List A
+boolFilter p = mapMaybe (λ x → if p x then just x else nothing)
-partition : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A)
-partition p [] = ([] , [])
-partition p (x ∷ xs) with p x | partition p xs
+boolPartition : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A)
+boolPartition p [] = ([] , [])
+boolPartition p (x ∷ xs) with p x | boolPartition p xs
... | true | (ys , zs) = (x ∷ ys , zs)
... | false | (ys , zs) = (ys , x ∷ zs)
+
+-- Version 0.16
+
+boolTakeWhile : ∀ {a} {A : Set a} → (A → Bool) → List A → List A
+boolTakeWhile p [] = []
+boolTakeWhile p (x ∷ xs) with p x
+... | true = x ∷ boolTakeWhile p xs
+... | false = []
+
+boolDropWhile : ∀ {a} {A : Set a} → (A → Bool) → List A → List A
+boolDropWhile p [] = []
+boolDropWhile p (x ∷ xs) with p x
+... | true = boolDropWhile p xs
+... | false = x ∷ xs
+
+boolSpan : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A)
+boolSpan p [] = ([] , [])
+boolSpan p (x ∷ xs) with p x
+... | true = Prod.map (x ∷_) id (boolSpan p xs)
+... | false = ([] , x ∷ xs)
+
+boolBreak : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A)
+boolBreak p = boolSpan (not ∘ p)
diff --git a/src/Data/List/Categorical.agda b/src/Data/List/Categorical.agda
new file mode 100644
index 0000000..1b6a07b
--- /dev/null
+++ b/src/Data/List/Categorical.agda
@@ -0,0 +1,252 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A categorical view of List
+------------------------------------------------------------------------
+
+module Data.List.Categorical where
+
+open import Category.Functor
+open import Category.Applicative
+open import Category.Monad
+open import Data.Bool.Base using (false; true)
+open import Data.List
+open import Data.List.Properties
+open import Function
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≢_; _≗_; refl)
+open P.≡-Reasoning
+
+------------------------------------------------------------------------
+-- List applicative functor
+
+functor : ∀ {ℓ} → RawFunctor {ℓ} List
+functor = record { _<$>_ = map }
+
+applicative : ∀ {ℓ} → RawApplicative {ℓ} List
+applicative = record
+ { pure = [_]
+ ; _⊛_ = λ fs as → concatMap (λ f → map f as) fs
+ }
+
+------------------------------------------------------------------------
+-- List monad
+
+monad : ∀ {ℓ} → RawMonad {ℓ} List
+monad = record
+ { return = [_]
+ ; _>>=_ = flip concatMap
+ }
+
+monadZero : ∀ {ℓ} → RawMonadZero {ℓ} List
+monadZero = record
+ { monad = monad
+ ; ∅ = []
+ }
+
+monadPlus : ∀ {ℓ} → RawMonadPlus {ℓ} List
+monadPlus = record
+ { monadZero = monadZero
+ ; _∣_ = _++_
+ }
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {f F} (App : RawApplicative {f} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → List (F A) → F (List A)
+ sequenceA [] = pure []
+ sequenceA (x ∷ xs) = _∷_ <$> x ⊛ sequenceA xs
+
+ mapA : ∀ {a} {A : Set a} {B} → (A → F B) → List A → F (List B)
+ mapA f = sequenceA ∘ map f
+
+ forA : ∀ {a} {A : Set a} {B} → List A → (A → F B) → F (List B)
+ forA = flip mapA
+
+module _ {m M} (Mon : RawMonad {m} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM = sequenceA App
+ mapM = mapA App
+ forM = forA App
+
+------------------------------------------------------------------------
+-- List monad transformer
+
+monadT : ∀ {ℓ} → RawMonadT {ℓ} (_∘′ List)
+monadT M = record
+ { return = pure ∘′ [_]
+ ; _>>=_ = λ mas f → mas >>= λ as → concat <$> mapM M f as
+ } where open RawMonad M
+
+------------------------------------------------------------------------
+-- The list monad.
+
+private
+ open module LMP {ℓ} = RawMonadPlus (monadPlus {ℓ = ℓ})
+
+module MonadProperties where
+
+ left-identity : ∀ {ℓ} {A B : Set ℓ} (x : A) (f : A → List B) →
+ (return x >>= f) ≡ f x
+ left-identity x f = ++-identityʳ (f x)
+
+ right-identity : ∀ {ℓ} {A : Set ℓ} (xs : List A) →
+ (xs >>= return) ≡ xs
+ right-identity [] = refl
+ right-identity (x ∷ xs) = P.cong (x ∷_) (right-identity xs)
+
+ left-zero : ∀ {ℓ} {A B : Set ℓ} (f : A → List B) → (∅ >>= f) ≡ ∅
+ left-zero f = refl
+
+ right-zero : ∀ {ℓ} {A B : Set ℓ} (xs : List A) →
+ (xs >>= const ∅) ≡ ∅ {A = B}
+ right-zero [] = refl
+ right-zero (x ∷ xs) = right-zero xs
+
+ private
+
+ not-left-distributive :
+ let xs = true ∷ false ∷ []; f = return; g = return in
+ (xs >>= λ x → f x ∣ g x) ≢ ((xs >>= f) ∣ (xs >>= g))
+ not-left-distributive ()
+
+ right-distributive : ∀ {ℓ} {A B : Set ℓ}
+ (xs ys : List A) (f : A → List B) →
+ (xs ∣ ys >>= f) ≡ ((xs >>= f) ∣ (ys >>= f))
+ right-distributive [] ys f = refl
+ right-distributive (x ∷ xs) ys f = begin
+ f x ∣ (xs ∣ ys >>= f) ≡⟨ P.cong (f x ∣_) $ right-distributive xs ys f ⟩
+ f x ∣ ((xs >>= f) ∣ (ys >>= f)) ≡⟨ P.sym $ ++-assoc (f x) _ _ ⟩
+ ((f x ∣ (xs >>= f)) ∣ (ys >>= f)) ∎
+
+ associative : ∀ {ℓ} {A B C : Set ℓ}
+ (xs : List A) (f : A → List B) (g : B → List C) →
+ (xs >>= λ x → f x >>= g) ≡ (xs >>= f >>= g)
+ associative [] f g = refl
+ associative (x ∷ xs) f g = begin
+ (f x >>= g) ∣ (xs >>= λ x → f x >>= g) ≡⟨ P.cong ((f x >>= g) ∣_) $ associative xs f g ⟩
+ (f x >>= g) ∣ (xs >>= f >>= g) ≡⟨ P.sym $ right-distributive (f x) (xs >>= f) g ⟩
+ (f x ∣ (xs >>= f) >>= g) ∎
+
+ cong : ∀ {ℓ} {A B : Set ℓ} {xs₁ xs₂} {f₁ f₂ : A → List B} →
+ xs₁ ≡ xs₂ → f₁ ≗ f₂ → (xs₁ >>= f₁) ≡ (xs₂ >>= f₂)
+ cong {xs₁ = xs} refl f₁≗f₂ = P.cong concat (map-cong f₁≗f₂ xs)
+
+------------------------------------------------------------------------
+-- The applicative functor derived from the list monad.
+
+-- Note that these proofs (almost) show that RawIMonad.rawIApplicative
+-- is correctly defined. The proofs can be reused if proof components
+-- are ever added to RawIMonad and RawIApplicative.
+
+module Applicative where
+
+ private
+
+ module MP = MonadProperties
+
+ -- A variant of flip map.
+
+ pam : ∀ {ℓ} {A B : Set ℓ} → List A → (A → B) → List B
+ pam xs f = xs >>= return ∘ f
+
+ -- ∅ is a left zero for _⊛_.
+
+ left-zero : ∀ {ℓ} {A B : Set ℓ} (xs : List A) → (∅ ⊛ xs) ≡ ∅ {A = B}
+ left-zero xs = begin
+ ∅ ⊛ xs ≡⟨⟩
+ (∅ >>= pam xs) ≡⟨ MonadProperties.left-zero (pam xs) ⟩
+ ∅ ∎
+
+ -- ∅ is a right zero for _⊛_.
+
+ right-zero : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) → (fs ⊛ ∅) ≡ ∅
+ right-zero {ℓ} fs = begin
+ fs ⊛ ∅ ≡⟨⟩
+ (fs >>= pam ∅) ≡⟨ (MP.cong (refl {x = fs}) λ f →
+ MP.left-zero (return ∘ f)) ⟩
+ (fs >>= λ _ → ∅) ≡⟨ MP.right-zero fs ⟩
+ ∅ ∎
+
+ -- _⊛_ distributes over _∣_ from the right.
+
+ right-distributive : ∀ {ℓ} {A B : Set ℓ} (fs₁ fs₂ : List (A → B)) xs →
+ ((fs₁ ∣ fs₂) ⊛ xs) ≡ (fs₁ ⊛ xs ∣ fs₂ ⊛ xs)
+ right-distributive fs₁ fs₂ xs = begin
+ (fs₁ ∣ fs₂) ⊛ xs ≡⟨⟩
+ (fs₁ ∣ fs₂ >>= pam xs) ≡⟨ MonadProperties.right-distributive fs₁ fs₂ (pam xs) ⟩
+ (fs₁ >>= pam xs) ∣ (fs₂ >>= pam xs) ≡⟨⟩
+ (fs₁ ⊛ xs ∣ fs₂ ⊛ xs) ∎
+
+ -- _⊛_ does not distribute over _∣_ from the left.
+
+ private
+
+ not-left-distributive :
+ let fs = id ∷ id ∷ []; xs₁ = true ∷ []; xs₂ = true ∷ false ∷ [] in
+ (fs ⊛ (xs₁ ∣ xs₂)) ≢ (fs ⊛ xs₁ ∣ fs ⊛ xs₂)
+ not-left-distributive ()
+
+ -- Applicative functor laws.
+
+ identity : ∀ {a} {A : Set a} (xs : List A) → (return id ⊛ xs) ≡ xs
+ identity xs = begin
+ return id ⊛ xs ≡⟨⟩
+ (return id >>= pam xs) ≡⟨ MonadProperties.left-identity id (pam xs) ⟩
+ (xs >>= return) ≡⟨ MonadProperties.right-identity xs ⟩
+ xs ∎
+
+ private
+
+ pam-lemma : ∀ {ℓ} {A B C : Set ℓ}
+ (xs : List A) (f : A → B) (fs : B → List C) →
+ (pam xs f >>= fs) ≡ (xs >>= λ x → fs (f x))
+ pam-lemma xs f fs = begin
+ (pam xs f >>= fs) ≡⟨ P.sym $ MP.associative xs (return ∘ f) fs ⟩
+ (xs >>= λ x → return (f x) >>= fs) ≡⟨ MP.cong (refl {x = xs}) (λ x → MP.left-identity (f x) fs) ⟩
+ (xs >>= λ x → fs (f x)) ∎
+
+ composition : ∀ {ℓ} {A B C : Set ℓ}
+ (fs : List (B → C)) (gs : List (A → B)) xs →
+ (return _∘′_ ⊛ fs ⊛ gs ⊛ xs) ≡ (fs ⊛ (gs ⊛ xs))
+ composition {ℓ} fs gs xs = begin
+ return _∘′_ ⊛ fs ⊛ gs ⊛ xs ≡⟨⟩
+ (return _∘′_ >>= pam fs >>= pam gs >>= pam xs) ≡⟨ MP.cong (MP.cong (MP.left-identity _∘′_ (pam fs))
+ (λ f → refl {x = pam gs f}))
+ (λ fg → refl {x = pam xs fg}) ⟩
+ (pam fs _∘′_ >>= pam gs >>= pam xs) ≡⟨ MP.cong (pam-lemma fs _∘′_ (pam gs)) (λ _ → refl) ⟩
+ ((fs >>= λ f → pam gs (f ∘′_)) >>= pam xs) ≡⟨ P.sym $ MP.associative fs (λ f → pam gs (_∘′_ f)) (pam xs) ⟩
+ (fs >>= λ f → pam gs (f ∘′_) >>= pam xs) ≡⟨ (MP.cong (refl {x = fs}) λ f →
+ pam-lemma gs (f ∘′_) (pam xs)) ⟩
+ (fs >>= λ f → gs >>= λ g → pam xs (f ∘′ g)) ≡⟨ (MP.cong (refl {x = fs}) λ f →
+ MP.cong (refl {x = gs}) λ g →
+ P.sym $ pam-lemma xs g (return ∘ f)) ⟩
+ (fs >>= λ f → gs >>= λ g → pam (pam xs g) f) ≡⟨ (MP.cong (refl {x = fs}) λ f →
+ MP.associative gs (pam xs) (return ∘ f)) ⟩
+ (fs >>= pam (gs >>= pam xs)) ≡⟨⟩
+ fs ⊛ (gs ⊛ xs) ∎
+
+ homomorphism : ∀ {ℓ} {A B : Set ℓ} (f : A → B) x →
+ (return f ⊛ return x) ≡ return (f x)
+ homomorphism f x = begin
+ return f ⊛ return x ≡⟨⟩
+ (return f >>= pam (return x)) ≡⟨ MP.left-identity f (pam (return x)) ⟩
+ pam (return x) f ≡⟨ MP.left-identity x (return ∘ f) ⟩
+ return (f x) ∎
+
+ interchange : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) {x} →
+ (fs ⊛ return x) ≡ (return (λ f → f x) ⊛ fs)
+ interchange fs {x} = begin
+ fs ⊛ return x ≡⟨⟩
+ (fs >>= pam (return x)) ≡⟨ (MP.cong (refl {x = fs}) λ f →
+ MP.left-identity x (return ∘ f)) ⟩
+ (fs >>= λ f → return (f x)) ≡⟨⟩
+ (pam fs (λ f → f x)) ≡⟨ P.sym $ MP.left-identity (λ f → f x) (pam fs) ⟩
+ (return (λ f → f x) >>= pam fs) ≡⟨⟩
+ return (λ f → f x) ⊛ fs ∎
diff --git a/src/Data/List/Countdown.agda b/src/Data/List/Countdown.agda
index 4818b07..4c93f7e 100644
--- a/src/Data/List/Countdown.agda
+++ b/src/Data/List/Countdown.agda
@@ -5,10 +5,10 @@
-- of elements /not/ in a given list
------------------------------------------------------------------------
-import Level
+open import Level using (0ℓ)
open import Relation.Binary
-module Data.List.Countdown (D : DecSetoid Level.zero Level.zero) where
+module Data.List.Countdown (D : DecSetoid 0ℓ 0ℓ) where
open import Data.Empty
open import Data.Fin using (Fin; zero; suc; punchOut)
@@ -18,11 +18,12 @@ open import Function
open import Function.Equality using (_⟨$⟩_)
open import Function.Injection
using (Injection; module Injection)
-open import Data.List
+open import Data.List hiding (lookup)
open import Data.List.Any as Any using (here; there)
open import Data.Nat.Base using (ℕ; zero; suc)
open import Data.Product
open import Data.Sum
+open import Data.Sum.Properties
open import Relation.Nullary
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; _≢_; refl; cong)
@@ -31,17 +32,13 @@ open PropEq.≡-Reasoning
private
open module D = DecSetoid D
hiding (refl) renaming (Carrier to Elem)
- open import Data.List.Any.Membership D.setoid
+ open import Data.List.Membership.Setoid D.setoid
------------------------------------------------------------------------
-- Helper functions
private
- drop-inj₂ : ∀ {A B : Set} {x y} →
- inj₂ {A = A} {B = B} x ≡ inj₂ y → x ≡ y
- drop-inj₂ refl = refl
-
-- The /first/ occurrence of x in xs.
first-occurrence : ∀ {xs} x → x ∈ xs → x ∈ xs
@@ -126,8 +123,8 @@ empty : ∀ {n} → Injection D.setoid (PropEq.setoid (Fin n)) → [] ⊕ n
empty inj =
record { kind = inj₂ ∘ _⟨$⟩_ to
; injective = λ {x} {y} {i} eq₁ eq₂ → injective (begin
- to ⟨$⟩ x ≡⟨ drop-inj₂ eq₁ ⟩
- i ≡⟨ PropEq.sym $ drop-inj₂ eq₂ ⟩
+ to ⟨$⟩ x ≡⟨ inj₂-injective eq₁ ⟩
+ i ≡⟨ PropEq.sym $ inj₂-injective eq₂ ⟩
to ⟨$⟩ y ∎)
}
where open Injection inj
@@ -199,7 +196,7 @@ insert {counted} {n} counted⊕1+n x x∉counted =
inj eq₁ eq₂ | no _ | no _ | inj₂ i | inj₂ _ | inj₂ _ | _ | _ | hlp =
hlp _ refl refl $
punchOut-injective {i = i} _ _ $
- (PropEq.trans (drop-inj₂ eq₁) (PropEq.sym (drop-inj₂ eq₂)))
+ (PropEq.trans (inj₂-injective eq₁) (PropEq.sym (inj₂-injective eq₂)))
-- Counts an element if it has not already been counted.
diff --git a/src/Data/List/Literals.agda b/src/Data/List/Literals.agda
new file mode 100644
index 0000000..9245077
--- /dev/null
+++ b/src/Data/List/Literals.agda
@@ -0,0 +1,19 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- List Literals
+------------------------------------------------------------------------
+
+module Data.List.Literals where
+
+open import Agda.Builtin.FromString
+open import Data.Unit
+open import Agda.Builtin.Char
+open import Agda.Builtin.List
+open import Data.String.Base using (toList)
+
+isString : IsString (List Char)
+isString = record
+ { Constraint = λ _ → ⊤
+ ; fromString = λ s → toList s
+ }
diff --git a/src/Data/List/Membership/DecPropositional.agda b/src/Data/List/Membership/DecPropositional.agda
new file mode 100644
index 0000000..5b59040
--- /dev/null
+++ b/src/Data/List/Membership/DecPropositional.agda
@@ -0,0 +1,19 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Decidable propositional membership over lists
+------------------------------------------------------------------------
+
+open import Relation.Binary using (Decidable)
+open import Relation.Binary.PropositionalEquality using (_≡_; decSetoid)
+
+module Data.List.Membership.DecPropositional
+ {a} {A : Set a} (_≟_ : Decidable (_≡_ {A = A})) where
+
+------------------------------------------------------------------------
+-- Re-export contents of propositional membership
+
+open import Data.List.Membership.Propositional public
+open import Data.List.Membership.DecSetoid (decSetoid _≟_) public
+ using (_∈?_)
+
diff --git a/src/Data/List/Membership/DecSetoid.agda b/src/Data/List/Membership/DecSetoid.agda
new file mode 100644
index 0000000..8ea8548
--- /dev/null
+++ b/src/Data/List/Membership/DecSetoid.agda
@@ -0,0 +1,25 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Decidable setoid membership over lists
+------------------------------------------------------------------------
+
+open import Relation.Binary using (Decidable; DecSetoid)
+
+module Data.List.Membership.DecSetoid {a ℓ} (DS : DecSetoid a ℓ) where
+
+open import Data.List.Any using (any)
+open DecSetoid DS
+
+------------------------------------------------------------------------
+-- Re-export contents of propositional membership
+
+open import Data.List.Membership.Setoid (DecSetoid.setoid DS) public
+
+------------------------------------------------------------------------
+-- Other operations
+
+infix 4 _∈?_
+
+_∈?_ : Decidable _∈_
+x ∈? xs = any (x ≟_) xs
diff --git a/src/Data/List/Membership/Propositional.agda b/src/Data/List/Membership/Propositional.agda
new file mode 100644
index 0000000..0de5e55
--- /dev/null
+++ b/src/Data/List/Membership/Propositional.agda
@@ -0,0 +1,24 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Data.List.Any.Membership instantiated with propositional equality,
+-- along with some additional definitions.
+------------------------------------------------------------------------
+
+module Data.List.Membership.Propositional {a} {A : Set a} where
+
+open import Data.List.Any using (Any)
+open import Relation.Binary.PropositionalEquality using (setoid; subst)
+
+import Data.List.Membership.Setoid as SetoidMembership
+
+------------------------------------------------------------------------
+-- Re-export contents of setoid membership
+
+open SetoidMembership (setoid A) public hiding (lose)
+
+------------------------------------------------------------------------
+-- Other operations
+
+lose : ∀ {p} {P : A → Set p} {x xs} → x ∈ xs → P x → Any P xs
+lose = SetoidMembership.lose (setoid A) (subst _)
diff --git a/src/Data/List/Membership/Propositional/Properties.agda b/src/Data/List/Membership/Propositional/Properties.agda
new file mode 100644
index 0000000..2422c22
--- /dev/null
+++ b/src/Data/List/Membership/Propositional/Properties.agda
@@ -0,0 +1,332 @@
+------------------------------------------------------------------------
+-- 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.Membership.Propositional.Properties where
+
+open import Algebra using (CommutativeSemiring)
+open import Algebra.FunctionProperties using (Op₂; Selective)
+open import Category.Monad using (RawMonad)
+open import Data.Bool.Base using (Bool; false; true; T)
+open import Data.Fin using (Fin)
+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.Membership.Propositional
+import Data.List.Membership.Setoid.Properties as Membershipₛ
+open import Data.List.Relation.Equality.Propositional
+ using (_≋_; ≡⇒≋; ≋⇒≡)
+open import Data.List.Categorical using (monad)
+open import Data.Nat using (ℕ; zero; suc; pred; s≤s; _≤_; _<_; _≤?_)
+open import Data.Nat.Properties
+open import Data.Product hiding (map)
+open import Data.Product.Relation.Pointwise.NonDependent using (_×-cong_)
+import Data.Product.Relation.Pointwise.Dependent as Σ
+open import Data.Sum as Sum using (_⊎_; inj₁; inj₂)
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence using (module Equivalence)
+open import Function.Injection using (Injection; Injective; _↣_)
+open import Function.Inverse as Inv using (_↔_; module Inverse)
+import Function.Related as Related
+open import Function.Related.TypeIsomorphisms
+open import Relation.Binary hiding (Decidable)
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≢_; refl; sym; trans; cong; subst; →-to-⟶; _≗_)
+import Relation.Binary.Properties.DecTotalOrder as DTOProperties
+open import Relation.Unary using (_⟨×⟩_; Decidable)
+open import Relation.Nullary using (¬_; Dec; yes; no)
+open import Relation.Nullary.Negation
+
+private
+ open module ListMonad {ℓ} = RawMonad (monad {ℓ = ℓ})
+
+------------------------------------------------------------------------
+-- Publicly re-export properties from Core
+
+open import Data.List.Membership.Propositional.Properties.Core public
+
+------------------------------------------------------------------------
+-- Equality
+
+module _ {a} {A : Set a} where
+
+ ∈-resp-≋ : ∀ {x} → (x ∈_) Respects _≋_
+ ∈-resp-≋ = Membershipₛ.∈-resp-≋ (P.setoid A)
+
+ ∉-resp-≋ : ∀ {x} → (x ∉_) Respects _≋_
+ ∉-resp-≋ = Membershipₛ.∉-resp-≋ (P.setoid A)
+
+------------------------------------------------------------------------
+-- mapWith∈
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ mapWith∈-cong : ∀ (xs : List A) → (f g : ∀ {x} → x ∈ xs → B) →
+ (∀ {x} → (x∈xs : x ∈ xs) → f x∈xs ≡ g x∈xs) →
+ mapWith∈ xs f ≡ mapWith∈ xs g
+ mapWith∈-cong [] f g cong = refl
+ mapWith∈-cong (x ∷ xs) f g cong = P.cong₂ _∷_ (cong (here refl))
+ (mapWith∈-cong xs (f ∘ there) (g ∘ there) (cong ∘ there))
+
+ mapWith∈≗map : ∀ f xs → mapWith∈ xs (λ {x} _ → f x) ≡ map f xs
+ mapWith∈≗map f xs =
+ ≋⇒≡ (Membershipₛ.mapWith∈≗map (P.setoid A) (P.setoid B) f xs)
+
+------------------------------------------------------------------------
+-- map
+
+module _ {a b} {A : Set a} {B : Set b} {f : A → B} where
+
+ ∈-map⁺ : ∀ {x xs} → x ∈ xs → f x ∈ map f xs
+ ∈-map⁺ = Membershipₛ.∈-map⁺ (P.setoid A) (P.setoid B) (P.cong f)
+
+ ∈-map⁻ : ∀ {y xs} → y ∈ map f xs → ∃ λ x → x ∈ xs × y ≡ f x
+ ∈-map⁻ = Membershipₛ.∈-map⁻ (P.setoid A) (P.setoid B)
+
+ map-∈↔ : ∀ {y xs} → (∃ λ x → x ∈ xs × y ≡ f x) ↔ y ∈ 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
+
+------------------------------------------------------------------------
+-- _++_
+
+module _ {a} (A : Set a) {v : A} where
+
+ ∈-++⁺ˡ : ∀ {xs ys} → v ∈ xs → v ∈ xs ++ ys
+ ∈-++⁺ˡ = Membershipₛ.∈-++⁺ˡ (P.setoid A)
+
+ ∈-++⁺ʳ : ∀ xs {ys} → v ∈ ys → v ∈ xs ++ ys
+ ∈-++⁺ʳ = Membershipₛ.∈-++⁺ʳ (P.setoid A)
+
+ ∈-++⁻ : ∀ xs {ys} → v ∈ xs ++ ys → (v ∈ xs) ⊎ (v ∈ ys)
+ ∈-++⁻ = Membershipₛ.∈-++⁻ (P.setoid A)
+
+ ∈-insert : ∀ xs {ys} → v ∈ xs ++ [ v ] ++ ys
+ ∈-insert xs = Membershipₛ.∈-insert (P.setoid A) xs refl
+
+ ∈-∃++ : ∀ {xs} → v ∈ xs → ∃₂ λ ys zs → xs ≡ ys ++ [ v ] ++ zs
+ ∈-∃++ v∈xs with Membershipₛ.∈-∃++ (P.setoid A) v∈xs
+ ... | ys , zs , _ , refl , eq = ys , zs , ≋⇒≡ eq
+
+------------------------------------------------------------------------
+-- concat
+
+module _ {a} {A : Set a} {v : A} where
+
+ ∈-concat⁺ : ∀ {xss} → Any (v ∈_) xss → v ∈ concat xss
+ ∈-concat⁺ = Membershipₛ.∈-concat⁺ (P.setoid A)
+
+ ∈-concat⁻ : ∀ xss → v ∈ concat xss → Any (v ∈_) xss
+ ∈-concat⁻ = Membershipₛ.∈-concat⁻ (P.setoid A)
+
+ ∈-concat⁺′ : ∀ {vs xss} → v ∈ vs → vs ∈ xss → v ∈ concat xss
+ ∈-concat⁺′ v∈vs vs∈xss =
+ Membershipₛ.∈-concat⁺′ (P.setoid A) v∈vs (Any.map ≡⇒≋ vs∈xss)
+
+ ∈-concat⁻′ : ∀ xss → v ∈ concat xss → ∃ λ xs → v ∈ xs × xs ∈ xss
+ ∈-concat⁻′ xss v∈c with Membershipₛ.∈-concat⁻′ (P.setoid A) xss v∈c
+ ... | xs , v∈xs , xs∈xss = xs , v∈xs , Any.map ≋⇒≡ xs∈xss
+
+ concat-∈↔ : ∀ {xss : List (List A)} →
+ (∃ λ xs → v ∈ xs × xs ∈ xss) ↔ v ∈ concat xss
+ concat-∈↔ {xss} =
+ (∃ λ xs → v ∈ xs × xs ∈ xss) ↔⟨ Σ.cong Inv.id $ ×-comm _ _ ⟩
+ (∃ λ xs → xs ∈ xss × v ∈ xs) ↔⟨ Any↔ ⟩
+ Any (Any (v ≡_)) xss ↔⟨ concat↔ ⟩
+ v ∈ concat xss ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- applyUpTo
+
+module _ {a} {A : Set a} where
+
+ ∈-applyUpTo⁺ : ∀ (f : ℕ → A) {i n} → i < n → f i ∈ applyUpTo f n
+ ∈-applyUpTo⁺ = Membershipₛ.∈-applyUpTo⁺ (P.setoid A)
+
+ ∈-applyUpTo⁻ : ∀ {v} f {n} → v ∈ applyUpTo f n →
+ ∃ λ i → i < n × v ≡ f i
+ ∈-applyUpTo⁻ = Membershipₛ.∈-applyUpTo⁻ (P.setoid A)
+
+------------------------------------------------------------------------
+-- tabulate
+
+module _ {a} {A : Set a} where
+
+ ∈-tabulate⁺ : ∀ {n} {f : Fin n → A} i → f i ∈ tabulate f
+ ∈-tabulate⁺ = Membershipₛ.∈-tabulate⁺ (P.setoid A)
+
+ ∈-tabulate⁻ : ∀ {n} {f : Fin n → A} {v} →
+ v ∈ tabulate f → ∃ λ i → v ≡ f i
+ ∈-tabulate⁻ = Membershipₛ.∈-tabulate⁻ (P.setoid A)
+
+------------------------------------------------------------------------
+-- filter
+
+module _ {a p} {A : Set a} {P : A → Set p} (P? : Decidable P) where
+
+ ∈-filter⁺ : ∀ {x xs} → x ∈ xs → P x → x ∈ filter P? xs
+ ∈-filter⁺ = Membershipₛ.∈-filter⁺ (P.setoid A) P? (P.subst P)
+
+ ∈-filter⁻ : ∀ {v xs} → v ∈ filter P? xs → v ∈ xs × P v
+ ∈-filter⁻ = Membershipₛ.∈-filter⁻ (P.setoid A) P? (P.subst P)
+
+------------------------------------------------------------------------
+-- _>>=_
+
+module _ {ℓ} {A B : Set ℓ} where
+
+ >>=-∈↔ : ∀ {xs} {f : A → List B} {y} →
+ (∃ λ x → x ∈ xs × y ∈ f x) ↔ y ∈ (xs >>= f)
+ >>=-∈↔ {xs = xs} {f} {y} =
+ (∃ λ x → x ∈ xs × y ∈ f x) ↔⟨ Any↔ ⟩
+ Any (Any (y ≡_) ∘ f) xs ↔⟨ >>=↔ ⟩
+ y ∈ (xs >>= f) ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- _⊛_
+
+module _ {ℓ} {A B : Set ℓ} where
+
+ ⊛-∈↔ : ∀ (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 Inv.id (∃∃↔∃∃ _) ⟩
+ (∃ λ f → f ∈ fs × ∃ λ x → x ∈ xs × y ≡ f x) ↔⟨ Σ.cong Inv.id ((_ ∎) ⟨ _×-cong_ ⟩ Any↔) ⟩
+ (∃ λ f → f ∈ fs × Any (_≡_ y ∘ f) xs) ↔⟨ Any↔ ⟩
+ Any (λ f → Any (_≡_ y ∘ f) xs) fs ↔⟨ ⊛↔ ⟩
+ y ∈ (fs ⊛ xs) ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- _⊗_
+
+module _ {ℓ} {A B : Set ℓ} where
+
+ ⊗-∈↔ : ∀ {xs ys} {x : A} {y : B} →
+ (x ∈ xs × y ∈ ys) ↔ (x , y) ∈ (xs ⊗ ys)
+ ⊗-∈↔ {xs} {ys} {x} {y} =
+ (x ∈ xs × y ∈ ys) ↔⟨ ⊗↔′ ⟩
+ Any (x ≡_ ⟨×⟩ y ≡_) (xs ⊗ ys) ↔⟨ Any-cong ×-≡×≡↔≡,≡ (_ ∎) ⟩
+ (x , y) ∈ (xs ⊗ ys) ∎
+ where
+ open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- length
+
+module _ {a} {A : Set a} where
+
+ ∈-length : ∀ {x xs} → x ∈ xs → 1 ≤ length xs
+ ∈-length = Membershipₛ.∈-length (P.setoid A)
+
+------------------------------------------------------------------------
+-- lookup
+
+module _ {a} {A : Set a} where
+
+ ∈-lookup : ∀ xs i → lookup xs i ∈ xs
+ ∈-lookup = Membershipₛ.∈-lookup (P.setoid A)
+
+------------------------------------------------------------------------
+-- foldr
+
+module _ {a} {A : Set a} {_•_ : Op₂ A} where
+
+ foldr-selective : Selective _≡_ _•_ → ∀ e xs →
+ (foldr _•_ e xs ≡ e) ⊎ (foldr _•_ e xs ∈ xs)
+ foldr-selective = Membershipₛ.foldr-selective (P.setoid A)
+
+------------------------------------------------------------------------
+-- Other properties
+
+-- Only a finite number of distinct elements can be members of a
+-- given list.
+
+module _ {a} {A : Set a} where
+
+ finite : (f : ℕ ↣ A) → ∀ xs → ¬ (∀ i → Injection.to f ⟨$⟩ i ∈ xs)
+ finite inj [] fᵢ∈[] = ¬Any[] (fᵢ∈[] 0)
+ finite inj (x ∷ xs) fᵢ∈x∷xs = excluded-middle helper
+ where
+ open Injection inj renaming (injective to f-inj)
+
+ f : ℕ → A
+ f = to ⟨$⟩_
+
+ not-x : ∀ {i} → f i ≢ x → f i ∈ xs
+ not-x {i} fᵢ≢x with fᵢ∈x∷xs i
+ ... | here fᵢ≡x = contradiction fᵢ≡x fᵢ≢x
+ ... | there fᵢ∈xs = fᵢ∈xs
+
+ helper : ¬ Dec (∃ λ i → f i ≡ x)
+ helper (no fᵢ≢x) = finite inj xs (λ i → not-x (fᵢ≢x ∘ _,_ i))
+ helper (yes (i , fᵢ≡x)) = finite f′-inj xs f′ⱼ∈xs
+ where
+ f′ : ℕ → A
+ f′ j with i ≤? j
+ ... | yes i≤j = f (suc j)
+ ... | no i≰j = f j
+
+ ∈-if-not-i : ∀ {j} → i ≢ j → f j ∈ xs
+ ∈-if-not-i i≢j = not-x (i≢j ∘ f-inj ∘ trans fᵢ≡x ∘ sym)
+
+ lemma : ∀ {k j} → i ≤ j → ¬ (i ≤ k) → suc j ≢ k
+ lemma i≤j i≰1+j refl = i≰1+j (≤-step i≤j)
+
+ f′ⱼ∈xs : ∀ j → f′ j ∈ xs
+ f′ⱼ∈xs j with i ≤? j
+ ... | yes i≤j = ∈-if-not-i (<⇒≢ (s≤s i≤j))
+ ... | no i≰j = ∈-if-not-i (<⇒≢ (≰⇒> i≰j) ∘ sym)
+
+ f′-injective′ : Injective {B = P.setoid A} (→-to-⟶ f′)
+ f′-injective′ {j} {k} eq with i ≤? j | i ≤? k
+ ... | yes _ | yes _ = P.cong pred (f-inj eq)
+ ... | yes i≤j | no i≰k = contradiction (f-inj eq) (lemma i≤j i≰k)
+ ... | no i≰j | yes i≤k = contradiction (f-inj eq) (lemma i≤k i≰j ∘ sym)
+ ... | no _ | no _ = f-inj eq
+
+ f′-inj = record
+ { to = →-to-⟶ f′
+ ; injective = f′-injective′
+ }
+
+------------------------------------------------------------------------
+-- DEPRECATED
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+boolFilter-∈ : ∀ {a} {A : Set a} (p : A → Bool) (xs : List A) {x} →
+ x ∈ xs → p x ≡ true → x ∈ boolFilter p xs
+boolFilter-∈ p [] () _
+boolFilter-∈ p (x ∷ xs) (here refl) px≡true rewrite px≡true = here refl
+boolFilter-∈ p (y ∷ xs) (there pxs) px≡true with p y
+... | true = there (boolFilter-∈ p xs pxs px≡true)
+... | false = boolFilter-∈ p xs pxs px≡true
+{-# WARNING_ON_USAGE boolFilter-∈
+"Warning: boolFilter was deprecated in v0.15.
+Please use filter instead."
+#-}
+
+-- Version 0.16
+
+filter-∈ = ∈-filter⁺
+{-# WARNING_ON_USAGE filter-∈
+"Warning: filter-∈ was deprecated in v0.16.
+Please use ∈-filter⁺ instead."
+#-}
+
diff --git a/src/Data/List/Membership/Propositional/Properties/Core.agda b/src/Data/List/Membership/Propositional/Properties/Core.agda
new file mode 100644
index 0000000..a2129c5
--- /dev/null
+++ b/src/Data/List/Membership/Propositional/Properties/Core.agda
@@ -0,0 +1,74 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Core properties related to propositional list membership.
+--
+-- This file is needed to break the cyclic dependency with the proof
+-- `Any-cong` in `Data.Any.Properties` which relies on `Any↔` in this
+-- file.
+------------------------------------------------------------------------
+
+module Data.List.Membership.Propositional.Properties.Core where
+
+open import Function using (flip; id; _∘_)
+open import Function.Inverse using (_↔_; inverse)
+open import Data.List.Base using (List)
+open import Data.List.Any as Any using (Any; here; there)
+open import Data.List.Membership.Propositional
+open import Data.Product as Prod
+ using (_,_; proj₁; proj₂; uncurry′; ∃; _×_)
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; refl)
+open import Relation.Unary using (_⊆_)
+
+-- Lemmas relating map and find.
+
+map∘find : ∀ {a p} {A : Set a} {P : A → Set p} {xs}
+ (p : Any P xs) → let p′ = find p in
+ {f : _≡_ (proj₁ p′) ⊆ P} →
+ f refl ≡ proj₂ (proj₂ p′) →
+ Any.map f (proj₁ (proj₂ p′)) ≡ p
+map∘find (here p) hyp = P.cong here hyp
+map∘find (there p) hyp = P.cong there (map∘find p hyp)
+
+find∘map : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q}
+ {xs : List A} (p : Any P xs) (f : P ⊆ Q) →
+ find (Any.map f p) ≡ Prod.map id (Prod.map id f) (find p)
+find∘map (here p) f = refl
+find∘map (there p) f rewrite find∘map p f = refl
+
+-- find satisfies a simple equality when the predicate is a
+-- propositional equality.
+
+find-∈ : ∀ {a} {A : Set a} {x : A} {xs : List A} (x∈xs : x ∈ xs) →
+ find x∈xs ≡ (x , x∈xs , refl)
+find-∈ (here refl) = refl
+find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl
+
+-- 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
+
+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↔ = inverse ∃∈-Any find from∘to lose∘find
+ where
+ from∘to : ∀ v → find (∃∈-Any v) ≡ v
+ from∘to p = find∘lose _ (proj₁ (proj₂ p)) (proj₂ (proj₂ p))
diff --git a/src/Data/List/Membership/Setoid.agda b/src/Data/List/Membership/Setoid.agda
new file mode 100644
index 0000000..a3384f1
--- /dev/null
+++ b/src/Data/List/Membership/Setoid.agda
@@ -0,0 +1,59 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- List membership and some related definitions
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Data.List.Membership.Setoid {c ℓ} (S : Setoid c ℓ) where
+
+open import Function using (_∘_; id; flip)
+open import Data.List.Base 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 Setoid S renaming (Carrier to A)
+
+------------------------------------------------------------------------
+-- Definitions
+
+infix 4 _∈_ _∉_
+
+_∈_ : A → List A → Set _
+x ∈ xs = Any (x ≈_) xs
+
+_∉_ : A → List A → Set _
+x ∉ xs = ¬ x ∈ xs
+
+------------------------------------------------------------------------
+-- Operations
+
+mapWith∈ : ∀ {b} {B : Set b}
+ (xs : List A) → (∀ {x} → x ∈ xs → B) → List B
+mapWith∈ [] f = []
+mapWith∈ (x ∷ xs) f = f (here refl) ∷ mapWith∈ xs (f ∘ there)
+
+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)
+
+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
+
+------------------------------------------------------------------------
+-- DEPRECATED
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.16
+
+map-with-∈ = mapWith∈
+{-# WARNING_ON_USAGE map-with-∈
+"Warning: map-with-∈ was deprecated in v0.16.
+Please use mapWith∈ instead."
+#-}
diff --git a/src/Data/List/Membership/Setoid/Properties.agda b/src/Data/List/Membership/Setoid/Properties.agda
new file mode 100644
index 0000000..a8f3a97
--- /dev/null
+++ b/src/Data/List/Membership/Setoid/Properties.agda
@@ -0,0 +1,239 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties related to setoid list membership
+------------------------------------------------------------------------
+
+module Data.List.Membership.Setoid.Properties where
+
+open import Algebra.FunctionProperties using (Op₂; Selective)
+open import Data.Fin using (Fin; zero; suc)
+open import Data.List
+open import Data.List.Any as Any using (Any; here; there)
+import Data.List.Any.Properties as Any
+import Data.List.Membership.Setoid as Membership
+import Data.List.Relation.Equality.Setoid as Equality
+open import Data.Nat using (z≤n; s≤s; _≤_; _<_)
+open import Data.Nat.Properties using (≤-trans; n≤1+n)
+open import Data.Product as Prod using (∃; _×_; _,_ ; ∃₂)
+open import Data.Sum using (_⊎_; inj₁; inj₂)
+open import Function using (flip; _∘_; id)
+open import Relation.Binary hiding (Decidable)
+open import Relation.Unary using (Decidable; Pred)
+open import Relation.Nullary using (yes; no)
+open import Relation.Nullary.Negation using (contradiction)
+open Setoid using (Carrier)
+
+------------------------------------------------------------------------
+-- Equality properties
+
+module _ {c ℓ} (S : Setoid c ℓ) where
+
+ open Setoid S
+ open Equality S
+ open Membership S
+
+ -- _∈_ respects the underlying equality
+
+ ∈-resp-≈ : ∀ {xs} → (_∈ xs) Respects _≈_
+ ∈-resp-≈ x≈y x∈xs = Any.map (trans (sym x≈y)) x∈xs
+
+ ∉-resp-≈ : ∀ {xs} → (_∉ xs) Respects _≈_
+ ∉-resp-≈ v≈w v∉xs w∈xs = v∉xs (∈-resp-≈ (sym v≈w) w∈xs)
+
+ ∈-resp-≋ : ∀ {x} → (x ∈_) Respects _≋_
+ ∈-resp-≋ = Any.lift-resp (flip trans)
+
+ ∉-resp-≋ : ∀ {x} → (x ∉_) Respects _≋_
+ ∉-resp-≋ xs≋ys v∉xs v∈ys = v∉xs (∈-resp-≋ (≋-sym xs≋ys) v∈ys)
+
+------------------------------------------------------------------------
+-- mapWith∈
+
+module _ {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 _≈₂_; refl to refl₂)
+ open Equality S₁ using ([]; _∷_) renaming (_≋_ to _≋₁_)
+ open Equality S₂ using () renaming (_≋_ to _≋₂_)
+ open Membership S₁
+
+ mapWith∈-cong : ∀ {xs ys} → xs ≋₁ ys →
+ (f : ∀ {x} → x ∈ xs → A₂) →
+ (g : ∀ {y} → y ∈ ys → A₂) →
+ (∀ {x y} → x ≈₁ y → (x∈xs : x ∈ xs) (y∈ys : y ∈ ys) →
+ f x∈xs ≈₂ g y∈ys) →
+ mapWith∈ xs f ≋₂ mapWith∈ ys g
+ mapWith∈-cong [] f g cong = []
+ mapWith∈-cong (x≈y ∷ xs≋ys) f g cong =
+ cong x≈y (here refl₁) (here refl₁) ∷
+ mapWith∈-cong xs≋ys (f ∘ there) (g ∘ there)
+ (λ x≈y x∈xs y∈ys → cong x≈y (there x∈xs) (there y∈ys))
+
+ mapWith∈≗map : ∀ f xs → mapWith∈ xs (λ {x} _ → f x) ≋₂ map f xs
+ mapWith∈≗map f [] = []
+ mapWith∈≗map f (x ∷ xs) = refl₂ ∷ mapWith∈≗map f xs
+
+------------------------------------------------------------------------
+-- map
+
+module _ {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 _≈₁_ ⟶ _≈₂_ → ∀ {v xs} →
+ v ∈₁ xs → f v ∈₂ map f xs
+ ∈-map⁺ pres x∈xs = Any.map⁺ (Any.map pres x∈xs)
+
+ ∈-map⁻ : ∀ {v xs f} → v ∈₂ map f xs →
+ ∃ λ x → x ∈₁ xs × v ≈₂ f x
+ ∈-map⁻ x∈map = find (Any.map⁻ x∈map)
+
+------------------------------------------------------------------------
+-- _++_
+
+module _ {c ℓ} (S : Setoid c ℓ) where
+
+ open Membership S using (_∈_)
+ open Setoid S
+ open Equality S using (_≋_; _∷_; ≋-refl)
+
+ ∈-++⁺ˡ : ∀ {v xs ys} → v ∈ xs → v ∈ xs ++ ys
+ ∈-++⁺ˡ = Any.++⁺ˡ
+
+ ∈-++⁺ʳ : ∀ {v} xs {ys} → v ∈ ys → v ∈ xs ++ ys
+ ∈-++⁺ʳ = Any.++⁺ʳ
+
+ ∈-++⁻ : ∀ {v} xs {ys} → v ∈ xs ++ ys → (v ∈ xs) ⊎ (v ∈ ys)
+ ∈-++⁻ = Any.++⁻
+
+ ∈-insert : ∀ xs {ys v w} → v ≈ w → v ∈ xs ++ [ w ] ++ ys
+ ∈-insert xs = Any.++-insert xs
+
+ ∈-∃++ : ∀ {v xs} → v ∈ xs → ∃₂ λ ys zs → ∃ λ w →
+ v ≈ w × xs ≋ ys ++ [ w ] ++ zs
+ ∈-∃++ (here px) = [] , _ , _ , px , ≋-refl
+ ∈-∃++ (there {d} v∈xs) with ∈-∃++ v∈xs
+ ... | hs , _ , _ , v≈v′ , eq = d ∷ hs , _ , _ , v≈v′ , refl ∷ eq
+
+------------------------------------------------------------------------
+-- concat
+
+module _ {c ℓ} (S : Setoid c ℓ) where
+
+ open Setoid S using (_≈_)
+ open Membership S using (_∈_)
+ open Equality S using (≋-setoid)
+ open Membership ≋-setoid using (find) renaming (_∈_ to _∈ₗ_)
+
+ ∈-concat⁺ : ∀ {v xss} → Any (v ∈_) xss → v ∈ concat xss
+ ∈-concat⁺ = Any.concat⁺
+
+ ∈-concat⁻ : ∀ {v} xss → v ∈ concat xss → Any (v ∈_) xss
+ ∈-concat⁻ = Any.concat⁻
+
+ ∈-concat⁺′ : ∀ {v vs xss} → v ∈ vs → vs ∈ₗ xss → v ∈ concat xss
+ ∈-concat⁺′ v∈vs = ∈-concat⁺ ∘ Any.map (flip (∈-resp-≋ S) v∈vs)
+
+ ∈-concat⁻′ : ∀ {v} xss → v ∈ concat xss → ∃ λ xs → v ∈ xs × xs ∈ₗ xss
+ ∈-concat⁻′ xss v∈c[xss] with find (∈-concat⁻ xss v∈c[xss])
+ ... | xs , t , s = xs , s , t
+
+------------------------------------------------------------------------
+-- applyUpTo
+
+module _ {c ℓ} (S : Setoid c ℓ) where
+
+ open Setoid S using (_≈_; refl)
+ open Membership S using (_∈_)
+
+ ∈-applyUpTo⁺ : ∀ f {i n} → i < n → f i ∈ applyUpTo f n
+ ∈-applyUpTo⁺ f = Any.applyUpTo⁺ f refl
+
+ ∈-applyUpTo⁻ : ∀ {v} f {n} → v ∈ applyUpTo f n →
+ ∃ λ i → i < n × v ≈ f i
+ ∈-applyUpTo⁻ = Any.applyUpTo⁻
+
+------------------------------------------------------------------------
+-- tabulate
+
+module _ {c ℓ} (S : Setoid c ℓ) where
+
+ open Setoid S using (_≈_; refl) renaming (Carrier to A)
+ open Membership S using (_∈_)
+
+ ∈-tabulate⁺ : ∀ {n} {f : Fin n → A} i → f i ∈ tabulate f
+ ∈-tabulate⁺ i = Any.tabulate⁺ i refl
+
+ ∈-tabulate⁻ : ∀ {n} {f : Fin n → A} {v} →
+ v ∈ tabulate f → ∃ λ i → v ≈ f i
+ ∈-tabulate⁻ = Any.tabulate⁻
+
+------------------------------------------------------------------------
+-- filter
+
+module _ {c ℓ p} (S : Setoid c ℓ) {P : Pred (Carrier S) p}
+ (P? : Decidable P) (resp : P Respects (Setoid._≈_ S)) where
+
+ open Setoid S using (_≈_; sym)
+ open Membership S using (_∈_)
+
+ ∈-filter⁺ : ∀ {v xs} → v ∈ xs → P v → v ∈ filter P? xs
+ ∈-filter⁺ {xs = x ∷ _} (here v≈x) Pv with P? x
+ ... | yes _ = here v≈x
+ ... | no ¬Px = contradiction (resp v≈x Pv) ¬Px
+ ∈-filter⁺ {xs = x ∷ _} (there v∈xs) Pv with P? x
+ ... | yes _ = there (∈-filter⁺ v∈xs Pv)
+ ... | no _ = ∈-filter⁺ v∈xs Pv
+
+ ∈-filter⁻ : ∀ {v xs} → v ∈ filter P? xs → v ∈ xs × P v
+ ∈-filter⁻ {xs = []} ()
+ ∈-filter⁻ {xs = x ∷ xs} v∈f[x∷xs] with P? x
+ ... | no _ = Prod.map there id (∈-filter⁻ v∈f[x∷xs])
+ ... | yes Px with v∈f[x∷xs]
+ ... | here v≈x = here v≈x , resp (sym v≈x) Px
+ ... | there v∈fxs = Prod.map there id (∈-filter⁻ v∈fxs)
+
+------------------------------------------------------------------------
+-- length
+
+module _ {c ℓ} (S : Setoid c ℓ) where
+
+ open Membership S using (_∈_)
+
+ ∈-length : ∀ {x xs} → x ∈ xs → 1 ≤ length xs
+ ∈-length (here px) = s≤s z≤n
+ ∈-length (there x∈xs) = ≤-trans (∈-length x∈xs) (n≤1+n _)
+
+------------------------------------------------------------------------
+-- lookup
+
+module _ {c ℓ} (S : Setoid c ℓ) where
+
+ open Setoid S using (refl)
+ open Membership S using (_∈_)
+
+ ∈-lookup : ∀ xs i → lookup xs i ∈ xs
+ ∈-lookup [] ()
+ ∈-lookup (x ∷ xs) zero = here refl
+ ∈-lookup (x ∷ xs) (suc i) = there (∈-lookup xs i)
+
+------------------------------------------------------------------------
+-- foldr
+
+module _ {c ℓ} (S : Setoid c ℓ) {_•_ : Op₂ (Carrier S)} where
+
+ open Setoid S using (_≈_; refl; sym; trans)
+ open Membership S using (_∈_)
+
+ foldr-selective : Selective _≈_ _•_ → ∀ e xs →
+ (foldr _•_ e xs ≈ e) ⊎ (foldr _•_ e xs ∈ xs)
+ foldr-selective •-sel i [] = inj₁ refl
+ foldr-selective •-sel i (x ∷ xs) with •-sel x (foldr _•_ i xs)
+ ... | inj₁ x•f≈x = inj₂ (here x•f≈x)
+ ... | inj₂ x•f≈f with foldr-selective •-sel i xs
+ ... | inj₁ f≈i = inj₁ (trans x•f≈f f≈i)
+ ... | inj₂ f∈xs = inj₂ (∈-resp-≈ S (sym x•f≈f) (there f∈xs))
diff --git a/src/Data/List/NonEmpty.agda b/src/Data/List/NonEmpty.agda
index 0d08bce..8733d10 100644
--- a/src/Data/List/NonEmpty.agda
+++ b/src/Data/List/NonEmpty.agda
@@ -10,9 +10,10 @@ open import Category.Monad
open import Data.Bool.Base using (Bool; false; true; not; T)
open import Data.Bool.Properties
open import Data.List as List using (List; []; _∷_)
-open import Data.Maybe.Base using (nothing; just)
+open import Data.Maybe.Base using (Maybe ; nothing; just)
open import Data.Nat as Nat
-open import Data.Product using (∃; proj₁; proj₂; _,_; ,_)
+open import Data.Product as Prod using (∃; _×_; proj₁; proj₂; _,_; -,_)
+open import Data.These as These using (These; this; that; these)
open import Data.Sum as Sum using (_⊎_; inj₁; inj₂)
open import Data.Unit
open import Data.Vec as Vec using (Vec; []; _∷_)
@@ -36,28 +37,41 @@ record List⁺ {a} (A : Set a) : Set a where
open List⁺ public
-[_] : ∀ {a} {A : Set a} → A → List⁺ A
-[ x ] = x ∷ []
+-- Basic combinators
-infixr 5 _∷⁺_
+module _ {a} {A : Set a} where
-_∷⁺_ : ∀ {a} {A : Set a} → A → List⁺ A → List⁺ A
-x ∷⁺ y ∷ xs = x ∷ y ∷ xs
+ uncons : List⁺ A → A × List A
+ uncons (hd ∷ tl) = hd , tl
-length : ∀ {a} {A : Set a} → List⁺ A → ℕ
-length (x ∷ xs) = suc (List.length xs)
+ [_] : A → List⁺ A
+ [ x ] = x ∷ []
+
+ infixr 5 _∷⁺_
+
+ _∷⁺_ : A → List⁺ A → List⁺ A
+ x ∷⁺ y ∷ xs = x ∷ y ∷ xs
+
+ length : List⁺ A → ℕ
+ length (x ∷ xs) = suc (List.length xs)
------------------------------------------------------------------------
-- Conversion
-toList : ∀ {a} {A : Set a} → List⁺ A → List A
-toList (x ∷ xs) = x ∷ xs
+module _ {a} {A : Set a} where
+
+ toList : List⁺ A → List A
+ toList (x ∷ xs) = x ∷ xs
+
+ fromList : List A → Maybe (List⁺ A)
+ fromList [] = nothing
+ fromList (x ∷ xs) = just (x ∷ xs)
-fromVec : ∀ {n a} {A : Set a} → Vec A (suc n) → List⁺ A
-fromVec (x ∷ xs) = x ∷ Vec.toList xs
+ fromVec : ∀ {n} → Vec A (suc n) → List⁺ A
+ fromVec (x ∷ xs) = x ∷ Vec.toList xs
-toVec : ∀ {a} {A : Set a} (xs : List⁺ A) → Vec A (length xs)
-toVec (x ∷ xs) = x ∷ Vec.fromList xs
+ toVec : (xs : List⁺ A) → Vec A (length xs)
+ toVec (x ∷ xs) = x ∷ Vec.fromList xs
lift : ∀ {a b} {A : Set a} {B : Set b} →
(∀ {m} → Vec A (suc m) → ∃ λ n → Vec B (suc n)) →
@@ -100,28 +114,63 @@ foldl₁ f = foldl f id
-- Append (several variants).
-infixr 5 _⁺++⁺_ _++⁺_ _⁺++_
+module _ {a} {A : Set a} where
-_⁺++⁺_ : ∀ {a} {A : Set a} → List⁺ A → List⁺ A → List⁺ A
-(x ∷ xs) ⁺++⁺ (y ∷ ys) = x ∷ (xs List.++ y ∷ ys)
+ infixr 5 _⁺++⁺_ _++⁺_ _⁺++_
-_⁺++_ : ∀ {a} {A : Set a} → List⁺ A → List A → List⁺ A
-(x ∷ xs) ⁺++ ys = x ∷ (xs List.++ ys)
+ _⁺++⁺_ : List⁺ A → List⁺ A → List⁺ A
+ (x ∷ xs) ⁺++⁺ (y ∷ ys) = x ∷ (xs List.++ y ∷ ys)
-_++⁺_ : ∀ {a} {A : Set a} → List A → List⁺ A → List⁺ A
-xs ++⁺ ys = List.foldr _∷⁺_ ys xs
+ _⁺++_ : List⁺ A → List A → List⁺ A
+ (x ∷ xs) ⁺++ ys = x ∷ (xs List.++ ys)
-concat : ∀ {a} {A : Set a} → List⁺ (List⁺ A) → List⁺ A
-concat (xs ∷ xss) = xs ⁺++ List.concat (List.map toList xss)
+ _++⁺_ : List A → List⁺ A → List⁺ A
+ xs ++⁺ ys = List.foldr _∷⁺_ ys xs
-monad : ∀ {f} → RawMonad (List⁺ {a = f})
-monad = record
- { return = [_]
- ; _>>=_ = λ xs f → concat (map f xs)
- }
+ concat : List⁺ (List⁺ A) → List⁺ A
+ concat (xs ∷ xss) = xs ⁺++ List.concat (List.map toList xss)
+
+concatMap : ∀ {a b} {A : Set a} {B : Set b} → (A → List⁺ B) → List⁺ A → List⁺ B
+concatMap f = concat ∘′ map f
+
+-- Reverse
reverse : ∀ {a} {A : Set a} → List⁺ A → List⁺ A
-reverse = lift (,_ ∘′ Vec.reverse)
+reverse = lift (-,_ ∘′ Vec.reverse)
+
+-- Align and Zip
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
+
+ alignWith : (These A B → C) → List⁺ A → List⁺ B → List⁺ C
+ alignWith f (a ∷ as) (b ∷ bs) = f (these a b) ∷ List.alignWith f as bs
+
+ zipWith : (A → B → C) → List⁺ A → List⁺ B → List⁺ C
+ zipWith f (a ∷ as) (b ∷ bs) = f a b ∷ List.zipWith f as bs
+
+ unalignWith : (A → These B C) → List⁺ A → These (List⁺ B) (List⁺ C)
+ unalignWith f = foldr (These.alignWith mcons mcons ∘′ f)
+ (These.map [_] [_] ∘′ f)
+
+ where mcons : ∀ {e} {E : Set e} → These E (List⁺ E) → List⁺ E
+ mcons = These.fold [_] id _∷⁺_
+
+ unzipWith : (A → B × C) → List⁺ A → List⁺ B × List⁺ C
+ unzipWith f (a ∷ as) = Prod.zip _∷_ _∷_ (f a) (List.unzipWith f as)
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ align : List⁺ A → List⁺ B → List⁺ (These A B)
+ align = alignWith id
+
+ zip : List⁺ A → List⁺ B → List⁺ (A × B)
+ zip = zipWith _,_
+
+ unalign : List⁺ (These A B) → These (List⁺ A) (List⁺ B)
+ unalign = unalignWith id
+
+ unzip : List⁺ (A × B) → List⁺ A × List⁺ B
+ unzip = unzipWith id
-- Snoc.
@@ -191,7 +240,7 @@ flatten-split p (x ∷ xs)
wordsBy : ∀ {a} {A : Set a} → (A → Bool) → List A → List (List⁺ A)
wordsBy p =
- List.gfilter Sum.[ const nothing , just ∘′ map proj₁ ] ∘ split p
+ List.mapMaybe Sum.[ const nothing , just ∘′ map proj₁ ] ∘ split p
------------------------------------------------------------------------
-- Examples
diff --git a/src/Data/List/NonEmpty/Categorical.agda b/src/Data/List/NonEmpty/Categorical.agda
new file mode 100644
index 0000000..65c0a3c
--- /dev/null
+++ b/src/Data/List/NonEmpty/Categorical.agda
@@ -0,0 +1,87 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A categorical view of List⁺
+------------------------------------------------------------------------
+
+module Data.List.NonEmpty.Categorical where
+
+open import Agda.Builtin.List
+import Data.List.Categorical as List
+open import Data.List.NonEmpty
+open import Data.Product using (uncurry)
+open import Category.Functor
+open import Category.Applicative
+open import Category.Monad
+open import Category.Comonad
+open import Function
+
+------------------------------------------------------------------------
+-- List⁺ applicative functor
+
+functor : ∀ {f} → RawFunctor {f} List⁺
+functor = record
+ { _<$>_ = map
+ }
+
+applicative : ∀ {f} → RawApplicative {f} List⁺
+applicative = record
+ { pure = [_]
+ ; _⊛_ = λ fs as → concatMap (λ f → map f as) fs
+ }
+
+------------------------------------------------------------------------
+-- List⁺ monad
+
+monad : ∀ {f} → RawMonad {f} List⁺
+monad = record
+ { return = [_]
+ ; _>>=_ = flip concatMap
+ }
+
+------------------------------------------------------------------------
+-- List⁺ comonad
+
+comonad : ∀ {f} → RawComonad {f} List⁺
+comonad = record
+ { extract = head
+ ; extend = λ f → uncurry (extend f) ∘′ uncons
+ } where
+
+ extend : ∀ {A B} → (List⁺ A → B) → A → List A → List⁺ B
+ extend f x xs@[] = f (x ∷ xs) ∷ []
+ extend f x xs@(y ∷ ys) = f (x ∷ xs) ∷⁺ extend f y ys
+
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {f F} (App : RawApplicative {f} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → List⁺ (F A) → F (List⁺ A)
+ sequenceA (x ∷ xs) = _∷_ <$> x ⊛ List.sequenceA App xs
+
+ mapA : ∀ {a} {A : Set a} {B} → (A → F B) → List⁺ A → F (List⁺ B)
+ mapA f = sequenceA ∘ map f
+
+ forA : ∀ {a} {A : Set a} {B} → List⁺ A → (A → F B) → F (List⁺ B)
+ forA = flip mapA
+
+module _ {m M} (Mon : RawMonad {m} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM = sequenceA App
+ mapM = mapA App
+ forM = forA App
+
+------------------------------------------------------------------------
+-- List⁺ monad transformer
+
+monadT : ∀ {f} → RawMonadT {f} (_∘′ List⁺)
+monadT M = record
+ { return = pure ∘′ [_]
+ ; _>>=_ = λ mas f → mas >>= λ as → concat <$> mapM M f as
+ } where open RawMonad M
diff --git a/src/Data/List/NonEmpty/Properties.agda b/src/Data/List/NonEmpty/Properties.agda
index 2d1d6cc..cd28716 100644
--- a/src/Data/List/NonEmpty/Properties.agda
+++ b/src/Data/List/NonEmpty/Properties.agda
@@ -8,6 +8,8 @@ module Data.List.NonEmpty.Properties where
open import Category.Monad
open import Data.List as List using (List; []; _∷_; _++_)
+open import Data.List.Categorical using () renaming (monad to listMonad)
+open import Data.List.NonEmpty.Categorical using () renaming (monad to list⁺Monad)
open import Data.List.NonEmpty as List⁺
open import Data.List.Properties
open import Function
@@ -16,10 +18,10 @@ open import Relation.Binary.PropositionalEquality
open ≡-Reasoning
private
open module LMo {a} =
- RawMonad {f = a} List.monad
+ RawMonad {f = a} listMonad
using () renaming (_>>=_ to _⋆>>=_)
open module L⁺Mo {a} =
- RawMonad {f = a} List⁺.monad
+ RawMonad {f = a} list⁺Monad
η : ∀ {a} {A : Set a}
(xs : List⁺ A) → head xs ∷ tail xs ≡ List⁺.toList xs
@@ -44,5 +46,7 @@ toList->>= : ∀ {ℓ} {A B : Set ℓ}
(List⁺.toList xs ⋆>>= List⁺.toList ∘ f) ≡
(List⁺.toList (xs >>= f))
toList->>= f (x ∷ xs) = begin
- List.concat (List.map (List⁺.toList ∘ f) (x ∷ xs)) ≡⟨ cong List.concat $ map-compose {g = List⁺.toList} {f = f} (x ∷ xs) ⟩
- List.concat (List.map List⁺.toList (List.map f (x ∷ xs))) ∎
+ List.concat (List.map (List⁺.toList ∘ f) (x ∷ xs))
+ ≡⟨ cong List.concat $ map-compose {g = List⁺.toList} (x ∷ xs) ⟩
+ List.concat (List.map List⁺.toList (List.map f (x ∷ xs)))
+ ∎
diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda
index 9fa36a7..ece0594 100644
--- a/src/Data/List/Properties.agda
+++ b/src/Data/List/Properties.agda
@@ -10,126 +10,350 @@
module Data.List.Properties where
open import Algebra
-import Algebra.Monoid-solver
-open import Category.Monad
+open import Algebra.Structures
+open import Algebra.FunctionProperties
open import Data.Bool.Base using (Bool; false; true; not; if_then_else_)
open import Data.List as List
open import Data.List.All using (All; []; _∷_)
+open import Data.List.Any using (Any; here; there)
open import Data.Maybe.Base using (Maybe; just; nothing)
open import Data.Nat
open import Data.Nat.Properties
-open import Data.Product as Prod hiding (map)
+open import Data.Fin using (Fin; zero; suc)
+open import Data.Product as Prod hiding (map; zip)
+open import Data.These as These using (These; this; that; these)
open import Function
-open import Algebra.FunctionProperties
import Relation.Binary.EqReasoning as EqR
open import Relation.Binary.PropositionalEquality as P
- using (_≡_; _≢_; _≗_; refl)
-open import Relation.Nullary using (yes; no)
+ using (_≡_; _≢_; _≗_; refl ; sym)
+open import Relation.Nullary using (¬_; yes; no)
+open import Relation.Nullary.Negation using (contradiction)
open import Relation.Nullary.Decidable using (⌊_⌋)
-open import Relation.Unary using (Decidable)
+open import Relation.Unary using (Pred; Decidable; ∁)
+open import Relation.Unary.Properties using (∁?)
+
+------------------------------------------------------------------------
+-- _∷_
+
+module _ {a} {A : Set a} {x y : A} {xs ys : List A} where
+
+ ∷-injective : x ∷ xs ≡ y List.∷ ys → x ≡ y × xs ≡ ys
+ ∷-injective refl = (refl , refl)
-private
- open module LMP {ℓ} = RawMonadPlus (List.monadPlus {ℓ = ℓ})
- module LM {a} {A : Set a} = Monoid (List.monoid A)
+ ∷-injectiveˡ : x ∷ xs ≡ y List.∷ ys → x ≡ y
+ ∷-injectiveˡ refl = refl
-module List-solver {a} {A : Set a} =
- Algebra.Monoid-solver (monoid A) renaming (id to nil)
+ ∷-injectiveʳ : x ∷ xs ≡ y List.∷ ys → xs ≡ ys
+ ∷-injectiveʳ refl = refl
------------------------------------------------------------------------
--- Equality
-
-∷-injective : ∀ {a} {A : Set a} {x y : A} {xs ys} →
- x ∷ xs ≡ y List.∷ ys → x ≡ y × xs ≡ ys
-∷-injective refl = (refl , refl)
-
-∷ʳ-injective : ∀ {a} {A : Set a} {x y : A} xs ys →
- xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys × x ≡ y
-∷ʳ-injective [] [] refl = (refl , refl)
-∷ʳ-injective (x ∷ xs) (y ∷ ys) eq with ∷-injective eq
-... | refl , eq′ = Prod.map (P.cong (x ∷_)) id (∷ʳ-injective xs ys eq′)
-∷ʳ-injective [] (_ ∷ []) ()
-∷ʳ-injective [] (_ ∷ _ ∷ _) ()
-∷ʳ-injective (_ ∷ []) [] ()
-∷ʳ-injective (_ ∷ _ ∷ _) [] ()
+-- map
+
+module _ {a} {A : Set a} where
+
+ map-id : map id ≗ id {A = List A}
+ map-id [] = refl
+ map-id (x ∷ xs) = P.cong (x ∷_) (map-id xs)
+
+ map-id₂ : ∀ {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)
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ map-++-commute : ∀ (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-cong : ∀ {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₂ : ∀ {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 : ∀ (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)
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
+
+ map-compose : {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)
------------------------------------------------------------------------
--- _++_
+-- mapMaybe
-right-identity-unique : ∀ {a} {A : Set a} (xs : List A) {ys} →
- xs ≡ xs ++ ys → ys ≡ []
-right-identity-unique [] refl = refl
-right-identity-unique (x ∷ xs) eq =
- right-identity-unique xs (proj₂ (∷-injective eq))
-
-left-identity-unique : ∀ {a} {A : Set a} {xs} (ys : List A) →
- xs ≡ ys ++ xs → ys ≡ []
-left-identity-unique [] _ = refl
-left-identity-unique {xs = []} (y ∷ ys) ()
-left-identity-unique {xs = x ∷ xs} (y ∷ ys) eq
- with left-identity-unique (ys ++ [ x ]) (begin
- xs ≡⟨ proj₂ (∷-injective eq) ⟩
- ys ++ x ∷ xs ≡⟨ P.sym (LM.assoc ys [ x ] xs) ⟩
- (ys ++ [ x ]) ++ xs ∎)
- where open P.≡-Reasoning
-left-identity-unique {xs = x ∷ xs} (y ∷ [] ) eq | ()
-left-identity-unique {xs = x ∷ xs} (y ∷ _ ∷ _) eq | ()
+module _ {a} {A : Set a} where
+
+ mapMaybe-just : (xs : List A) → mapMaybe just xs ≡ xs
+ mapMaybe-just [] = refl
+ mapMaybe-just (x ∷ xs) = P.cong (x ∷_) (mapMaybe-just xs)
+
+ mapMaybe-nothing : (xs : List A) →
+ mapMaybe {B = A} (λ _ → nothing) xs ≡ []
+ mapMaybe-nothing [] = refl
+ mapMaybe-nothing (x ∷ xs) = mapMaybe-nothing 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)
+module _ {a b} {A : Set a} {B : Set b} (f : A → Maybe B) where
+
+ mapMaybe-concatMap : mapMaybe f ≗ concatMap (fromMaybe ∘ f)
+ mapMaybe-concatMap [] = refl
+ mapMaybe-concatMap (x ∷ xs) with f x
+ ... | just y = P.cong (y ∷_) (mapMaybe-concatMap xs)
+ ... | nothing = mapMaybe-concatMap xs
+
+ length-mapMaybe : ∀ xs → length (mapMaybe f xs) ≤ length xs
+ length-mapMaybe [] = z≤n
+ length-mapMaybe (x ∷ xs) with f x
+ ... | just y = s≤s (length-mapMaybe xs)
+ ... | nothing = ≤-step (length-mapMaybe xs)
------------------------------------------------------------------------
--- map
+-- _++_
+
+module _ {a} {A : Set a} where
-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)
+ ++-assoc : Associative {A = List A} _≡_ _++_
+ ++-assoc [] ys zs = refl
+ ++-assoc (x ∷ xs) ys zs = P.cong (x ∷_) (++-assoc xs ys zs)
-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)
+ ++-identityˡ : LeftIdentity {A = List A} _≡_ [] _++_
+ ++-identityˡ xs = refl
-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)
+ ++-identityʳ : RightIdentity {A = List A} _≡_ [] _++_
+ ++-identityʳ [] = refl
+ ++-identityʳ (x ∷ xs) = P.cong (x ∷_) (++-identityʳ 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)
+ ++-identity : Identity {A = List A} _≡_ [] _++_
+ ++-identity = ++-identityˡ , ++-identityʳ
-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)
+ ++-identityʳ-unique : ∀ (xs : List A) {ys} → xs ≡ xs ++ ys → ys ≡ []
+ ++-identityʳ-unique [] refl = refl
+ ++-identityʳ-unique (x ∷ xs) eq =
+ ++-identityʳ-unique xs (proj₂ (∷-injective eq))
-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)
+ ++-identityˡ-unique : ∀ {xs} (ys : List A) → xs ≡ ys ++ xs → ys ≡ []
+ ++-identityˡ-unique [] _ = refl
+ ++-identityˡ-unique {xs = []} (y ∷ ys) ()
+ ++-identityˡ-unique {xs = x ∷ xs} (y ∷ ys) eq
+ with ++-identityˡ-unique (ys ++ [ x ]) (begin
+ xs ≡⟨ proj₂ (∷-injective eq) ⟩
+ ys ++ x ∷ xs ≡⟨ P.sym (++-assoc ys [ x ] xs) ⟩
+ (ys ++ [ x ]) ++ xs ∎)
+ where open P.≡-Reasoning
+ ++-identityˡ-unique {xs = x ∷ xs} (y ∷ [] ) eq | ()
+ ++-identityˡ-unique {xs = x ∷ xs} (y ∷ _ ∷ _) eq | ()
+
+ length-++ : ∀ (xs : List A) {ys} → length (xs ++ ys) ≡ length xs + length ys
+ length-++ [] = refl
+ length-++ (x ∷ xs) = P.cong suc (length-++ xs)
+
+ ++-isSemigroup : IsSemigroup {A = List A} _≡_ _++_
+ ++-isSemigroup = record
+ { isEquivalence = P.isEquivalence
+ ; assoc = ++-assoc
+ ; ∙-cong = P.cong₂ _++_
+ }
+
+ ++-isMonoid : IsMonoid {A = List A} _≡_ _++_ []
+ ++-isMonoid = record
+ { isSemigroup = ++-isSemigroup
+ ; identity = ++-identity
+ }
+
+++-semigroup : ∀ {a} (A : Set a) → Semigroup _ _
+++-semigroup A = record
+ { Carrier = List A
+ ; _≈_ = _≡_
+ ; _∙_ = _++_
+ ; isSemigroup = ++-isSemigroup
+ }
+
+++-monoid : ∀ {a} (A : Set a) → Monoid _ _
+++-monoid A = record
+ { Carrier = List A
+ ; _≈_ = _≡_
+ ; _∙_ = _++_
+ ; ε = []
+ ; isMonoid = ++-isMonoid
+ }
-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)
+------------------------------------------------------------------------
+-- alignWith
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c}
+ {f g : These A B → C} where
+
+ alignWith-cong : f ≗ g → ∀ as → alignWith f as ≗ alignWith g as
+ alignWith-cong f≗g [] bs = map-cong (f≗g ∘ that) bs
+ alignWith-cong f≗g as@(_ ∷ _) [] = map-cong (f≗g ∘ this) as
+ alignWith-cong f≗g (a ∷ as) (b ∷ bs) =
+ P.cong₂ _∷_ (f≗g (these a b)) (alignWith-cong f≗g as bs)
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c}
+ (f : These A B → C) where
+
+ length-alignWith : ∀ xs ys →
+ length (alignWith f xs ys) ≡ length xs ⊔ length ys
+ length-alignWith [] ys = length-map (f ∘′ that) ys
+ length-alignWith xs@(_ ∷ _) [] = length-map (f ∘′ this) xs
+ length-alignWith (x ∷ xs) (y ∷ ys) = P.cong suc (length-alignWith xs ys)
+
+ alignWith-map : ∀ {d e} {D : Set d} {E : Set e} (g : D → A) (h : E → B) →
+ ∀ xs ys → alignWith f (map g xs) (map h ys) ≡
+ alignWith (f ∘′ These.map g h) xs ys
+ alignWith-map g h [] ys = sym (map-compose ys)
+ alignWith-map g h xs@(_ ∷ _) [] = sym (map-compose xs)
+ alignWith-map g h (x ∷ xs) (y ∷ ys) =
+ P.cong₂ _∷_ refl (alignWith-map g h xs ys)
+
+ map-alignWith : ∀ {d} {D : Set d} (g : C → D) → ∀ xs ys →
+ map g (alignWith f xs ys) ≡
+ alignWith (g ∘′ f) xs ys
+ map-alignWith g [] ys = sym (map-compose ys)
+ map-alignWith g xs@(_ ∷ _) [] = sym (map-compose xs)
+ map-alignWith g (x ∷ xs) (y ∷ ys) =
+ P.cong₂ _∷_ refl (map-alignWith g xs ys)
------------------------------------------------------------------------
--- replicate
+-- zipWith
+
+module _ {a b} {A : Set a} {B : Set b} (f : A → A → B) where
+
+ zipWith-comm : (∀ x y → f x y ≡ f y x) →
+ ∀ xs ys → zipWith f xs ys ≡ zipWith f ys xs
+ zipWith-comm f-comm [] [] = refl
+ zipWith-comm f-comm [] (x ∷ ys) = refl
+ zipWith-comm f-comm (x ∷ xs) [] = refl
+ zipWith-comm f-comm (x ∷ xs) (y ∷ ys) =
+ P.cong₂ _∷_ (f-comm x y) (zipWith-comm f-comm xs ys)
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c}
+ (f : A → B → C) where
+
+ zipWith-identityˡ : ∀ xs → zipWith f [] xs ≡ []
+ zipWith-identityˡ [] = refl
+ zipWith-identityˡ (x ∷ xs) = refl
+
+ zipWith-identityʳ : ∀ xs → zipWith f xs [] ≡ []
+ zipWith-identityʳ [] = refl
+ zipWith-identityʳ (x ∷ xs) = refl
+
+ length-zipWith : ∀ xs ys →
+ length (zipWith f xs ys) ≡ length xs ⊓ length ys
+ length-zipWith [] [] = refl
+ length-zipWith [] (y ∷ ys) = refl
+ length-zipWith (x ∷ xs) [] = refl
+ length-zipWith (x ∷ xs) (y ∷ ys) = P.cong suc (length-zipWith xs ys)
+
+ zipWith-map : ∀ {d e} {D : Set d} {E : Set e} (g : D → A) (h : E → B) →
+ ∀ xs ys → zipWith f (map g xs) (map h ys) ≡
+ zipWith (λ x y → f (g x) (h y)) xs ys
+ zipWith-map g h [] [] = refl
+ zipWith-map g h [] (y ∷ ys) = refl
+ zipWith-map g h (x ∷ xs) [] = refl
+ zipWith-map g h (x ∷ xs) (y ∷ ys) =
+ P.cong₂ _∷_ refl (zipWith-map g h xs ys)
+
+ map-zipWith : ∀ {d} {D : Set d} (g : C → D) → ∀ xs ys →
+ map g (zipWith f xs ys) ≡
+ zipWith (λ x y → g (f x y)) xs ys
+ map-zipWith g [] [] = refl
+ map-zipWith g [] (y ∷ ys) = refl
+ map-zipWith g (x ∷ xs) [] = refl
+ map-zipWith g (x ∷ xs) (y ∷ ys) =
+ P.cong₂ _∷_ refl (map-zipWith g xs ys)
-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)
+------------------------------------------------------------------------
+-- unalignWith
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ unalignWith-this : unalignWith ((A → These A B) ∋ this) ≗ (_, [])
+ unalignWith-this [] = refl
+ unalignWith-this (a ∷ as) = P.cong (Prod.map₁ (a ∷_)) (unalignWith-this as)
+
+ unalignWith-that : unalignWith ((B → These A B) ∋ that) ≗ ([] ,_)
+ unalignWith-that [] = refl
+ unalignWith-that (b ∷ bs) = P.cong (Prod.map₂ (b ∷_)) (unalignWith-that bs)
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c}
+ {f g : C → These A B} where
+
+ unalignWith-cong : f ≗ g → unalignWith f ≗ unalignWith g
+ unalignWith-cong f≗g [] = refl
+ unalignWith-cong f≗g (c ∷ cs) with f c | g c | f≗g c
+ ... | this a | ._ | refl = P.cong (Prod.map₁ (a ∷_)) (unalignWith-cong f≗g cs)
+ ... | that b | ._ | refl = P.cong (Prod.map₂ (b ∷_)) (unalignWith-cong f≗g cs)
+ ... | these a b | ._ | refl = P.cong (Prod.map (a ∷_) (b ∷_)) (unalignWith-cong f≗g cs)
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c}
+ (f : C → These A B) where
+
+ unalignWith-map : ∀ {d} {D : Set d} (g : D → C) →
+ ∀ ds → unalignWith f (map g ds) ≡ unalignWith (f ∘′ g) ds
+ unalignWith-map g [] = refl
+ unalignWith-map g (d ∷ ds) with f (g d)
+ ... | this a = P.cong (Prod.map₁ (a ∷_)) (unalignWith-map g ds)
+ ... | that b = P.cong (Prod.map₂ (b ∷_)) (unalignWith-map g ds)
+ ... | these a b = P.cong (Prod.map (a ∷_) (b ∷_)) (unalignWith-map g ds)
+
+ map-unalignWith : ∀ {d e} {D : Set d} {E : Set e} (g : A → D) (h : B → E) →
+ Prod.map (map g) (map h) ∘′ unalignWith f ≗ unalignWith (These.map g h ∘′ f)
+ map-unalignWith g h [] = refl
+ map-unalignWith g h (c ∷ cs) with f c
+ ... | this a = P.cong (Prod.map₁ (g a ∷_)) (map-unalignWith g h cs)
+ ... | that b = P.cong (Prod.map₂ (h b ∷_)) (map-unalignWith g h cs)
+ ... | these a b = P.cong (Prod.map (g a ∷_) (h b ∷_)) (map-unalignWith g h cs)
+
+ unalignWith-alignWith : (g : These A B → C) → f ∘′ g ≗ id →
+ ∀ as bs → unalignWith f (alignWith g as bs) ≡ (as , bs)
+ unalignWith-alignWith g g∘f≗id [] bs = begin
+ unalignWith f (map (g ∘′ that) bs) ≡⟨ unalignWith-map (g ∘′ that) bs ⟩
+ unalignWith (f ∘′ g ∘′ that) bs ≡⟨ unalignWith-cong (g∘f≗id ∘ that) bs ⟩
+ unalignWith that bs ≡⟨ unalignWith-that bs ⟩
+ [] , bs ∎ where open P.≡-Reasoning
+ unalignWith-alignWith g g∘f≗id as@(_ ∷ _) [] = begin
+ unalignWith f (map (g ∘′ this) as) ≡⟨ unalignWith-map (g ∘′ this) as ⟩
+ unalignWith (f ∘′ g ∘′ this) as ≡⟨ unalignWith-cong (g∘f≗id ∘ this) as ⟩
+ unalignWith this as ≡⟨ unalignWith-this as ⟩
+ as , [] ∎ where open P.≡-Reasoning
+ unalignWith-alignWith g g∘f≗id (a ∷ as) (b ∷ bs)
+ rewrite g∘f≗id (these a b) = let ih = unalignWith-alignWith g g∘f≗id as bs in
+ P.cong (Prod.map (a ∷_) (b ∷_)) ih
+
+------------------------------------------------------------------------
+-- unzipWith
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c}
+ (f : A → B × C) where
+
+ length-unzipWith₁ : ∀ xys →
+ length (proj₁ (unzipWith f xys)) ≡ length xys
+ length-unzipWith₁ [] = refl
+ length-unzipWith₁ (x ∷ xys) = P.cong suc (length-unzipWith₁ xys)
+
+ length-unzipWith₂ : ∀ xys →
+ length (proj₂ (unzipWith f xys)) ≡ length xys
+ length-unzipWith₂ [] = refl
+ length-unzipWith₂ (x ∷ xys) = P.cong suc (length-unzipWith₂ xys)
+
+ zipWith-unzipWith : (g : B → C → A) → uncurry′ g ∘ f ≗ id →
+ uncurry′ (zipWith g) ∘ (unzipWith f) ≗ id
+ zipWith-unzipWith g f∘g≗id [] = refl
+ zipWith-unzipWith g f∘g≗id (x ∷ xs) =
+ P.cong₂ _∷_ (f∘g≗id x) (zipWith-unzipWith g f∘g≗id xs)
------------------------------------------------------------------------
-- foldr
foldr-universal : ∀ {a b} {A : Set a} {B : Set b}
- (h : List A → B) f e →
- (h [] ≡ e) →
+ (h : List A → B) f e → (h [] ≡ e) →
(∀ x xs → h (x ∷ xs) ≡ f x (h xs)) →
h ≗ foldr f e
foldr-universal h f e base step [] = base
@@ -143,7 +367,7 @@ foldr-universal h f e base step (x ∷ xs) = begin
where open P.≡-Reasoning
foldr-cong : ∀ {a b} {A : Set a} {B : Set b}
- {f g : A → B → B} {d e : 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
@@ -157,19 +381,19 @@ foldr-fusion h {f} {g} e fuse =
foldr-universal (h ∘ foldr f e) g (h e) refl
(λ x xs → fuse x (foldr f e xs))
-idIsFold : ∀ {a} {A : Set a} → id {A = List A} ≗ foldr _∷_ []
-idIsFold = foldr-universal id _∷_ [] refl (λ _ _ → refl)
+id-is-foldr : ∀ {a} {A : Set a} → id {A = List A} ≗ foldr _∷_ []
+id-is-foldr = foldr-universal id _∷_ [] refl (λ _ _ → refl)
-++IsFold : ∀ {a} {A : Set a} (xs ys : List A) →
+++-is-foldr : ∀ {a} {A : Set a} (xs ys : List A) →
xs ++ ys ≡ foldr _∷_ ys xs
-++IsFold xs ys =
+++-is-foldr xs ys =
begin
xs ++ ys
- ≡⟨ P.cong (_++ ys) (idIsFold xs) ⟩
+ ≡⟨ P.cong (_++ ys) (id-is-foldr xs) ⟩
foldr _∷_ [] xs ++ ys
≡⟨ foldr-fusion (_++ ys) [] (λ _ _ → refl) xs ⟩
foldr _∷_ ([] ++ ys) xs
- ≡⟨ refl ⟩
+ ≡⟨⟩
foldr _∷_ ys xs
where open P.≡-Reasoning
@@ -179,12 +403,12 @@ foldr-++ : ∀ {a b} {A : Set a} {B : Set b} (f : A → B → B) x ys zs →
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-is-foldr : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} →
map f ≗ foldr (λ x ys → f x ∷ ys) []
-mapIsFold {f = f} =
+map-is-foldr {f = f} =
begin
map f
- ≈⟨ P.cong (map f) ∘ idIsFold ⟩
+ ≈⟨ P.cong (map f) ∘ id-is-foldr ⟩
map f ∘ foldr _∷_ []
≈⟨ foldr-fusion (map f) [] (λ _ _ → refl) ⟩
foldr (λ x ys → f x ∷ ys) []
@@ -199,32 +423,35 @@ 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
+module _ {a b} {A : Set a} {B : Set b} where
+
+ foldl-++ : ∀ (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
+ foldl-∷ʳ : ∀ (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} =
- begin
- concat ∘ map (map f)
- ≈⟨ P.cong concat ∘ mapIsFold {b = b} ⟩
- concat ∘ foldr (λ xs ys → map f xs ∷ ys) []
- ≈⟨ foldr-fusion {b = b} concat [] (λ _ _ → refl) ⟩
- foldr (λ ys zs → map f ys ++ zs) []
- ≈⟨ P.sym ∘ foldr-fusion (map f) [] (map-++-commute f) ⟩
- map f ∘ concat
- ∎
- where open EqR (P._→-setoid_ _ _)
+module _ {a b} {A : Set a} {B : Set b} where
+
+ concat-map : ∀ {f : A → B} → concat ∘ map (map f) ≗ map f ∘ concat
+ concat-map {f = f} =
+ begin
+ concat ∘ map (map f)
+ ≈⟨ P.cong concat ∘ map-is-foldr ⟩
+ concat ∘ foldr (λ xs → map f xs ∷_) []
+ ≈⟨ foldr-fusion concat [] (λ _ _ → refl) ⟩
+ foldr (λ ys → map f ys ++_) []
+ ≈⟨ P.sym ∘ foldr-fusion (map f) [] (map-++-commute f) ⟩
+ map f ∘ concat
+ ∎
+ where open EqR (P._→-setoid_ _ _)
------------------------------------------------------------------------
-- sum
@@ -233,359 +460,383 @@ 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) ≡⟨ P.sym (+-assoc x _ _) ⟩
(x + sum xs) + sum ys ∎
where open P.≡-Reasoning
------------------------------------------------------------------------
--- 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 (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 : ∀ {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)
-... | 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)
-... | false = refl
+-- replicate
+
+module _ {a} {A : Set a} where
+
+ length-replicate : ∀ n {x : A} → length (replicate n x) ≡ n
+ length-replicate zero = refl
+ length-replicate (suc n) = P.cong suc (length-replicate n)
------------------------------------------------------------------------
--- 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
-... | 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-nothing : ∀ {a} {A : Set a} (xs : List A) →
- 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-concatMap f [] = refl
-gfilter-concatMap f (x ∷ xs) with f x
-... | just y = P.cong (y ∷_) (gfilter-concatMap f xs)
-... | nothing = gfilter-concatMap f xs
-
-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
+-- scanr
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ scanr-defn : ∀ (f : A → B → B) (e : B) →
+ scanr f e ≗ map (foldr f e) ∘ tails
+ scanr-defn f e [] = refl
+ scanr-defn f e (x ∷ []) = refl
+ scanr-defn f e (x ∷ y ∷ xs)
+ with scanr f e (y ∷ xs) | scanr-defn f e (y ∷ xs)
+ ... | [] | ()
+ ... | z ∷ zs | eq with ∷-injective eq
+ ... | z≡fy⦇f⦈xs , _ = P.cong₂ (λ z → f x z ∷_) z≡fy⦇f⦈xs eq
------------------------------------------------------------------------
--- Inits, tails, and scanr
-
-scanr-defn : ∀ {a b} {A : Set a} {B : Set b}
- (f : A → B → B) (e : B) →
- scanr f e ≗ map (foldr f e) ∘ tails
-scanr-defn f e [] = refl
-scanr-defn f e (x ∷ []) = refl
-scanr-defn f e (x₁ ∷ x₂ ∷ xs)
- with scanr f e (x₂ ∷ xs) | scanr-defn f e (x₂ ∷ xs)
-... | [] | ()
-... | y ∷ ys | eq with ∷-injective eq
-... | y≡fx₂⦇f⦈xs , _ = P.cong₂ (λ z zs → f x₁ z ∷ zs) y≡fx₂⦇f⦈xs eq
-
-scanl-defn : ∀ {a b} {A : Set a} {B : Set b}
- (f : A → B → A) (e : A) →
- scanl f e ≗ map (foldl f e) ∘ inits
-scanl-defn f e [] = refl
-scanl-defn f e (x ∷ xs) = P.cong (e ∷_) (begin
+-- scanl
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ scanl-defn : ∀ (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 f (f e x) xs
- ≡⟨ scanl-defn f (f e x) xs ⟩
+ ≡⟨ scanl-defn f (f e x) xs ⟩
map (foldl f (f e x)) (inits xs)
- ≡⟨ refl ⟩
+ ≡⟨ refl ⟩
map (foldl f e ∘ (x ∷_)) (inits xs)
- ≡⟨ map-compose (inits xs) ⟩
+ ≡⟨ map-compose (inits xs) ⟩
map (foldl f e) (map (x ∷_) (inits xs))
- ∎)
- where open P.≡-Reasoning
+ ∎)
+ where open P.≡-Reasoning
------------------------------------------------------------------------
--- reverse
+-- tabulate
-unfold-reverse : ∀ {a} {A : Set a} (x : A) (xs : List A) →
- reverse (x ∷ xs) ≡ reverse xs ∷ʳ x
-unfold-reverse {A = A} x xs = helper [ x ] xs
- where
- open P.≡-Reasoning
- 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 (_++ 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 [] 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 (_++ [ x ]) (reverse-++-commute xs ys) ⟩
- (reverse ys ++ reverse xs) ++ [ x ] ≡⟨ LM.assoc (reverse ys) _ _ ⟩
- reverse ys ++ (reverse xs ++ [ x ]) ≡⟨ P.sym $ P.cong (reverse ys ++_) (unfold-reverse x xs) ⟩
- reverse ys ++ reverse (x ∷ xs) ∎
- where open P.≡-Reasoning
+module _ {a} {A : Set a} where
-reverse-map-commute :
- ∀ {a b} {A : Set a} {B : Set b} (f : A → B) → (xs : List A) →
- map f (reverse xs) ≡ reverse (map f xs)
-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 (_∷ʳ 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
+ tabulate-cong : ∀ {n} {f g : Fin n → A} →
+ f ≗ g → tabulate f ≡ tabulate g
+ tabulate-cong {zero} p = P.refl
+ tabulate-cong {suc n} p = P.cong₂ _∷_ (p zero) (tabulate-cong (p ∘ suc))
-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 (x ∷_) $ reverse-involutive xs ⟩
- x ∷ xs ∎
- where open P.≡-Reasoning
+ tabulate-lookup : ∀ (xs : List A) → tabulate (lookup xs) ≡ xs
+ tabulate-lookup [] = refl
+ tabulate-lookup (x ∷ xs) = P.cong (_ ∷_) (tabulate-lookup xs)
-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
+------------------------------------------------------------------------
+-- take
-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
+module _ {a} {A : Set a} where
-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
+ length-take : ∀ n (xs : List A) → length (take n xs) ≡ n ⊓ (length xs)
+ length-take zero xs = refl
+ length-take (suc n) [] = refl
+ length-take (suc n) (x ∷ xs) = P.cong suc (length-take n xs)
+
+------------------------------------------------------------------------
+-- drop
+
+module _ {a} {A : Set a} where
+
+ length-drop : ∀ n (xs : List A) → length (drop n xs) ≡ length xs ∸ n
+ length-drop zero xs = refl
+ length-drop (suc n) [] = refl
+ length-drop (suc n) (x ∷ xs) = length-drop n xs
+
+ take++drop : ∀ 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 (x ∷_) (take++drop n xs)
+
+------------------------------------------------------------------------
+-- splitAt
+
+module _ {a} {A : Set a} where
+
+ splitAt-defn : ∀ 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
------------------------------------------------------------------------
--- The list monad.
+-- takeWhile, dropWhile, and span
-module Monad where
+module _ {a p} {A : Set a} {P : Pred A p} (P? : Decidable P) where
- left-zero : ∀ {ℓ} {A B : Set ℓ} (f : A → List B) → (∅ >>= f) ≡ ∅
- left-zero f = refl
+ takeWhile++dropWhile : ∀ xs → takeWhile P? xs ++ dropWhile P? xs ≡ xs
+ takeWhile++dropWhile [] = refl
+ takeWhile++dropWhile (x ∷ xs) with P? x
+ ... | yes _ = P.cong (x ∷_) (takeWhile++dropWhile xs)
+ ... | no _ = refl
- right-zero : ∀ {ℓ} {A B : Set ℓ} (xs : List A) →
- (xs >>= const ∅) ≡ ∅ {A = B}
- right-zero [] = refl
- right-zero (x ∷ xs) = right-zero xs
+ span-defn : span P? ≗ < takeWhile P? , dropWhile P? >
+ span-defn [] = refl
+ span-defn (x ∷ xs) with P? x
+ ... | yes _ = P.cong (Prod.map (x ∷_) id) (span-defn xs)
+ ... | no _ = refl
- private
+------------------------------------------------------------------------
+-- filter
+
+module _ {a p} {A : Set a} {P : A → Set p} (P? : Decidable P) where
+
+ length-filter : ∀ xs → length (filter P? xs) ≤ length xs
+ length-filter [] = z≤n
+ length-filter (x ∷ xs) with P? x
+ ... | no _ = ≤-step (length-filter xs)
+ ... | yes _ = s≤s (length-filter xs)
+
+ filter-all : ∀ {xs} → All P xs → filter P? xs ≡ xs
+ filter-all {[]} [] = refl
+ filter-all {x ∷ xs} (px ∷ pxs) with P? x
+ ... | no ¬px = contradiction px ¬px
+ ... | yes _ = P.cong (x ∷_) (filter-all pxs)
+
+ filter-notAll : ∀ xs → Any (∁ P) xs → length (filter P? xs) < length xs
+ filter-notAll [] ()
+ filter-notAll (x ∷ xs) (here ¬px) with P? x
+ ... | no _ = s≤s (length-filter xs)
+ ... | yes px = contradiction px ¬px
+ filter-notAll (x ∷ xs) (there any) with P? x
+ ... | no _ = ≤-step (filter-notAll xs any)
+ ... | yes _ = s≤s (filter-notAll xs any)
+
+ filter-some : ∀ {xs} → Any P xs → 0 < length (filter P? xs)
+ filter-some {x ∷ xs} (here px) with P? x
+ ... | yes _ = s≤s z≤n
+ ... | no ¬px = contradiction px ¬px
+ filter-some {x ∷ xs} (there pxs) with P? x
+ ... | yes _ = ≤-step (filter-some pxs)
+ ... | no _ = filter-some pxs
+
+ filter-none : ∀ {xs} → All (∁ P) xs → filter P? xs ≡ []
+ filter-none {[]} [] = refl
+ filter-none {x ∷ xs} (¬px ∷ ¬pxs) with P? x
+ ... | no _ = filter-none ¬pxs
+ ... | yes px = contradiction px ¬px
+
+ filter-complete : ∀ {xs} → length (filter P? xs) ≡ length xs →
+ filter P? xs ≡ xs
+ filter-complete {[]} eq = refl
+ filter-complete {x ∷ xs} eq with P? x
+ ... | no ¬px = contradiction eq (<⇒≢ (s≤s (length-filter xs)))
+ ... | yes px = P.cong (x ∷_) (filter-complete (suc-injective eq))
+
+------------------------------------------------------------------------
+-- partition
+
+module _ {a p} {A : Set a} {P : A → Set p} (P? : Decidable P) where
+
+ partition-defn : partition P? ≗ < filter P? , filter (∁? P?) >
+ partition-defn [] = refl
+ partition-defn (x ∷ xs) with P? x
+ ... | yes Px = P.cong (Prod.map (x ∷_) id) (partition-defn xs)
+ ... | no ¬Px = P.cong (Prod.map id (x ∷_)) (partition-defn xs)
+
+------------------------------------------------------------------------
+-- reverse
+
+module _ {a} {A : Set a} where
+
+ unfold-reverse : ∀ (x : A) xs → reverse (x ∷ xs) ≡ reverse xs ∷ʳ x
+ unfold-reverse x xs = helper [ x ] xs
+ where
+ open P.≡-Reasoning
+ 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 (++-assoc (reverse ys) _ _) ⟩
+ (reverse ys ∷ʳ y) ++ xs ≡⟨ P.sym $ P.cong (_++ xs) (unfold-reverse y ys) ⟩
+ reverse (y ∷ ys) ++ xs ∎
+
+ reverse-++-commute : (xs ys : List A) →
+ reverse (xs ++ ys) ≡ reverse ys ++ reverse xs
+ reverse-++-commute [] ys = P.sym (++-identityʳ _)
+ reverse-++-commute (x ∷ xs) ys = begin
+ reverse (x ∷ xs ++ ys) ≡⟨ unfold-reverse x (xs ++ ys) ⟩
+ reverse (xs ++ ys) ++ [ x ] ≡⟨ P.cong (_++ [ x ]) (reverse-++-commute xs ys) ⟩
+ (reverse ys ++ reverse xs) ++ [ x ] ≡⟨ ++-assoc (reverse ys) _ _ ⟩
+ reverse ys ++ (reverse xs ++ [ x ]) ≡⟨ P.sym $ P.cong (reverse ys ++_) (unfold-reverse x xs) ⟩
+ reverse ys ++ reverse (x ∷ xs) ∎
+ where open P.≡-Reasoning
- not-left-distributive :
- let xs = true ∷ false ∷ []; f = return; g = return in
- (xs >>= λ x → f x ∣ g x) ≢ ((xs >>= f) ∣ (xs >>= g))
- not-left-distributive ()
+ reverse-involutive : 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 (x ∷_) $ reverse-involutive xs ⟩
+ x ∷ xs ∎
+ where open P.≡-Reasoning
- right-distributive : ∀ {ℓ} {A B : Set ℓ}
- (xs ys : List A) (f : A → List B) →
- (xs ∣ ys >>= f) ≡ ((xs >>= f) ∣ (ys >>= f))
- right-distributive [] ys f = refl
- right-distributive (x ∷ xs) ys f = begin
- f x ∣ (xs ∣ ys >>= f) ≡⟨ P.cong (_∣_ (f x)) $ right-distributive xs ys f ⟩
- f x ∣ ((xs >>= f) ∣ (ys >>= f)) ≡⟨ P.sym $ LM.assoc (f x) _ _ ⟩
- ((f x ∣ (xs >>= f)) ∣ (ys >>= f)) ∎
+ length-reverse : (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
- left-identity : ∀ {ℓ} {A B : Set ℓ} (x : A) (f : A → List B) →
- (return x >>= f) ≡ f x
- left-identity {ℓ} x f = proj₂ (LM.identity {a = ℓ}) (f x)
-
- right-identity : ∀ {a} {A : Set a} (xs : List A) →
- (xs >>= return) ≡ xs
- right-identity [] = refl
- right-identity (x ∷ xs) = P.cong (_∷_ x) (right-identity xs)
-
- associative : ∀ {ℓ} {A B C : Set ℓ}
- (xs : List A) (f : A → List B) (g : B → List C) →
- (xs >>= λ x → f x >>= g) ≡ (xs >>= f >>= g)
- associative [] f g = refl
- associative (x ∷ xs) f g = begin
- (f x >>= g) ∣ (xs >>= λ x → f x >>= g) ≡⟨ P.cong (_∣_ (f x >>= g)) $ associative xs f g ⟩
- (f x >>= g) ∣ (xs >>= f >>= g) ≡⟨ P.sym $ right-distributive (f x) (xs >>= f) g ⟩
- (f x ∣ (xs >>= f) >>= g) ∎
+module _ {a b} {A : Set a} {B : Set b} where
+
+ reverse-map-commute : (f : A → B) (xs : List A) →
+ map f (reverse xs) ≡ reverse (map f xs)
+ 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 (_∷ʳ 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
- cong : ∀ {ℓ} {A B : Set ℓ} {xs₁ xs₂} {f₁ f₂ : A → List B} →
- xs₁ ≡ xs₂ → f₁ ≗ f₂ → (xs₁ >>= f₁) ≡ (xs₂ >>= f₂)
- cong {xs₁ = xs} refl f₁≗f₂ = P.cong concat (map-cong f₁≗f₂ xs)
-
--- The applicative functor derived from the list monad.
-
--- Note that these proofs (almost) show that RawIMonad.rawIApplicative
--- is correctly defined. The proofs can be reused if proof components
--- are ever added to RawIMonad and RawIApplicative.
-
-module Applicative where
-
- open P.≡-Reasoning
-
- private
-
- -- A variant of flip map.
-
- pam : ∀ {ℓ} {A B : Set ℓ} → List A → (A → B) → List B
- pam xs f = xs >>= return ∘ f
-
- -- ∅ is a left zero for _⊛_.
-
- left-zero : ∀ {ℓ} {A B : Set ℓ} (xs : List A) → (∅ ⊛ xs) ≡ ∅ {A = B}
- left-zero xs = begin
- ∅ ⊛ xs ≡⟨ refl ⟩
- (∅ >>= pam xs) ≡⟨ Monad.left-zero (pam xs) ⟩
- ∅ ∎
-
- -- ∅ is a right zero for _⊛_.
-
- right-zero : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) → (fs ⊛ ∅) ≡ ∅
- right-zero {ℓ} fs = begin
- fs ⊛ ∅ ≡⟨ refl ⟩
- (fs >>= pam ∅) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
- Monad.left-zero (return {ℓ = ℓ} ∘ f)) ⟩
- (fs >>= λ _ → ∅) ≡⟨ Monad.right-zero fs ⟩
- ∅ ∎
-
- -- _⊛_ distributes over _∣_ from the right.
-
- right-distributive :
- ∀ {ℓ} {A B : Set ℓ} (fs₁ fs₂ : List (A → B)) xs →
- ((fs₁ ∣ fs₂) ⊛ xs) ≡ (fs₁ ⊛ xs ∣ fs₂ ⊛ xs)
- right-distributive fs₁ fs₂ xs = begin
- (fs₁ ∣ fs₂) ⊛ xs ≡⟨ refl ⟩
- (fs₁ ∣ fs₂ >>= pam xs) ≡⟨ Monad.right-distributive fs₁ fs₂ (pam xs) ⟩
- (fs₁ >>= pam xs) ∣ (fs₂ >>= pam xs) ≡⟨ refl ⟩
- (fs₁ ⊛ xs ∣ fs₂ ⊛ xs) ∎
-
- -- _⊛_ does not distribute over _∣_ from the left.
-
- private
-
- not-left-distributive :
- let fs = id ∷ id ∷ []; xs₁ = true ∷ []; xs₂ = true ∷ false ∷ [] in
- (fs ⊛ (xs₁ ∣ xs₂)) ≢ (fs ⊛ xs₁ ∣ fs ⊛ xs₂)
- not-left-distributive ()
-
- -- Applicative functor laws.
-
- identity : ∀ {a} {A : Set a} (xs : List A) → (return id ⊛ xs) ≡ xs
- identity xs = begin
- return id ⊛ xs ≡⟨ refl ⟩
- (return id >>= pam xs) ≡⟨ Monad.left-identity id (pam xs) ⟩
- (xs >>= return) ≡⟨ Monad.right-identity xs ⟩
- xs ∎
-
- private
-
- pam-lemma : ∀ {ℓ} {A B C : Set ℓ}
- (xs : List A) (f : A → B) (fs : B → List C) →
- (pam xs f >>= fs) ≡ (xs >>= λ x → fs (f x))
- pam-lemma xs f fs = begin
- (pam xs f >>= fs) ≡⟨ P.sym $ Monad.associative xs (return ∘ f) fs ⟩
- (xs >>= λ x → return (f x) >>= fs) ≡⟨ Monad.cong (refl {x = xs}) (λ x → Monad.left-identity (f x) fs) ⟩
- (xs >>= λ x → fs (f x)) ∎
-
- composition :
- ∀ {ℓ} {A B C : Set ℓ}
- (fs : List (B → C)) (gs : List (A → B)) xs →
- (return _∘′_ ⊛ fs ⊛ gs ⊛ xs) ≡ (fs ⊛ (gs ⊛ xs))
- composition {ℓ} fs gs xs = begin
- return _∘′_ ⊛ fs ⊛ gs ⊛ xs ≡⟨ refl ⟩
- (return _∘′_ >>= pam fs >>= pam gs >>= pam xs) ≡⟨ Monad.cong (Monad.cong (Monad.left-identity _∘′_ (pam fs))
- (λ f → refl {x = pam gs f}))
- (λ fg → refl {x = pam xs fg}) ⟩
- (pam fs _∘′_ >>= pam gs >>= pam xs) ≡⟨ Monad.cong (pam-lemma fs _∘′_ (pam gs)) (λ _ → refl) ⟩
- ((fs >>= λ f → pam gs (_∘′_ f)) >>= pam xs) ≡⟨ P.sym $ Monad.associative fs (λ f → pam gs (_∘′_ f)) (pam xs) ⟩
- (fs >>= λ f → pam gs (_∘′_ f) >>= pam xs) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
- pam-lemma gs (_∘′_ f) (pam xs)) ⟩
- (fs >>= λ f → gs >>= λ g → pam xs (f ∘′ g)) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
- Monad.cong (refl {x = gs}) λ g →
- P.sym $ pam-lemma xs g (return ∘ f)) ⟩
- (fs >>= λ f → gs >>= λ g → pam (pam xs g) f) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
- Monad.associative gs (pam xs) (return ∘ f)) ⟩
- (fs >>= pam (gs >>= pam xs)) ≡⟨ refl ⟩
- fs ⊛ (gs ⊛ xs) ∎
-
- homomorphism : ∀ {ℓ} {A B : Set ℓ} (f : A → B) x →
- (return f ⊛ return x) ≡ return (f x)
- homomorphism f x = begin
- return f ⊛ return x ≡⟨ refl ⟩
- (return f >>= pam (return x)) ≡⟨ Monad.left-identity f (pam (return x)) ⟩
- pam (return x) f ≡⟨ Monad.left-identity x (return ∘ f) ⟩
- return (f x) ∎
-
- interchange : ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) {x} →
- (fs ⊛ return x) ≡ (return (λ f → f x) ⊛ fs)
- interchange fs {x} = begin
- fs ⊛ return x ≡⟨ refl ⟩
- (fs >>= pam (return x)) ≡⟨ (Monad.cong (refl {x = fs}) λ f →
- Monad.left-identity x (return ∘ f)) ⟩
- (fs >>= λ f → return (f x)) ≡⟨ refl ⟩
- (pam fs (λ f → f x)) ≡⟨ P.sym $ Monad.left-identity (λ f → f x) (pam fs) ⟩
- (return (λ f → f x) >>= pam fs) ≡⟨ refl ⟩
- return (λ f → f x) ⊛ fs ∎
+ reverse-foldr : ∀ (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 : ∀ (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
+
+------------------------------------------------------------------------
+-- _∷ʳ_
+
+module _ {a} {A : Set a} where
+
+ ∷ʳ-injective : ∀ {x y : A} xs ys →
+ xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys × x ≡ y
+ ∷ʳ-injective [] [] refl = (refl , refl)
+ ∷ʳ-injective (x ∷ xs) (y ∷ ys) eq with ∷-injective eq
+ ... | refl , eq′ = Prod.map (P.cong (x ∷_)) id (∷ʳ-injective xs ys eq′)
+ ∷ʳ-injective [] (_ ∷ []) ()
+ ∷ʳ-injective [] (_ ∷ _ ∷ _) ()
+ ∷ʳ-injective (_ ∷ []) [] ()
+ ∷ʳ-injective (_ ∷ _ ∷ _) [] ()
+
+ ∷ʳ-injectiveˡ : ∀ {x y : A} (xs ys : List A) → xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys
+ ∷ʳ-injectiveˡ xs ys eq = proj₁ (∷ʳ-injective xs ys eq)
+
+ ∷ʳ-injectiveʳ : ∀ {x y : A} (xs ys : List A) → xs ∷ʳ x ≡ ys ∷ʳ y → x ≡ y
+ ∷ʳ-injectiveʳ xs ys eq = proj₂ (∷ʳ-injective xs ys eq)
+
+------------------------------------------------------------------------
+-- DEPRECATED
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+gfilter-just = mapMaybe-just
+{-# WARNING_ON_USAGE gfilter-just
+"Warning: gfilter-just was deprecated in v0.15.
+Please use mapMaybe-just instead."
+#-}
+gfilter-nothing = mapMaybe-nothing
+{-# WARNING_ON_USAGE gfilter-nothing
+"Warning: gfilter-nothing was deprecated in v0.15.
+Please use mapMaybe-nothing instead."
+#-}
+gfilter-concatMap = mapMaybe-concatMap
+{-# WARNING_ON_USAGE gfilter-concatMap
+"Warning: gfilter-concatMap was deprecated in v0.15.
+Please use mapMaybe-concatMap instead."
+#-}
+length-gfilter = length-mapMaybe
+{-# WARNING_ON_USAGE length-gfilter
+"Warning: length-gfilter was deprecated in v0.15.
+Please use length-mapMaybe instead."
+#-}
+right-identity-unique = ++-identityʳ-unique
+{-# WARNING_ON_USAGE right-identity-unique
+"Warning: right-identity-unique was deprecated in v0.15.
+Please use ++-identityʳ-unique instead."
+#-}
+left-identity-unique = ++-identityˡ-unique
+{-# WARNING_ON_USAGE left-identity-unique
+"Warning: left-identity-unique was deprecated in v0.15.
+Please use ++-identityˡ-unique instead."
+#-}
+
+-- Version 0.16
+
+module _ {a} {A : Set a} (p : A → Bool) where
+
+ boolTakeWhile++boolDropWhile : ∀ xs → boolTakeWhile p xs ++ boolDropWhile p xs ≡ xs
+ boolTakeWhile++boolDropWhile [] = refl
+ boolTakeWhile++boolDropWhile (x ∷ xs) with p x
+ ... | true = P.cong (x ∷_) (boolTakeWhile++boolDropWhile xs)
+ ... | false = refl
+ {-# WARNING_ON_USAGE boolTakeWhile++boolDropWhile
+ "Warning: boolTakeWhile and boolDropWhile were deprecated in v0.16.
+ Please use takeWhile and dropWhile instead."
+ #-}
+ boolSpan-defn : boolSpan p ≗ < boolTakeWhile p , boolDropWhile p >
+ boolSpan-defn [] = refl
+ boolSpan-defn (x ∷ xs) with p x
+ ... | true = P.cong (Prod.map (x ∷_) id) (boolSpan-defn xs)
+ ... | false = refl
+ {-# WARNING_ON_USAGE boolSpan-defn
+ "Warning: boolSpan, boolTakeWhile and boolDropWhile were deprecated in v0.16.
+ Please use span, takeWhile and dropWhile instead."
+ #-}
+ length-boolFilter : ∀ xs → length (boolFilter p xs) ≤ length xs
+ length-boolFilter xs =
+ length-mapMaybe (λ x → if p x then just x else nothing) xs
+ {-# WARNING_ON_USAGE length-boolFilter
+ "Warning: boolFilter was deprecated in v0.16.
+ Please use filter instead."
+ #-}
+ boolPartition-defn : boolPartition p ≗ < boolFilter p , boolFilter (not ∘ p) >
+ boolPartition-defn [] = refl
+ boolPartition-defn (x ∷ xs) with p x
+ ... | true = P.cong (Prod.map (x ∷_) id) (boolPartition-defn xs)
+ ... | false = P.cong (Prod.map id (x ∷_)) (boolPartition-defn xs)
+ {-# WARNING_ON_USAGE boolPartition-defn
+ "Warning: boolPartition and boolFilter were deprecated in v0.16.
+ Please use partition and filter instead."
+ #-}
+
+module _ {a p} {A : Set a} (P : A → Set p) (P? : Decidable P) where
+
+ boolFilter-filters : ∀ xs → All P (boolFilter (⌊_⌋ ∘ P?) xs)
+ boolFilter-filters [] = []
+ boolFilter-filters (x ∷ xs) with P? x
+ ... | yes px = px ∷ boolFilter-filters xs
+ ... | no ¬px = boolFilter-filters xs
+ {-# WARNING_ON_USAGE boolFilter-filters
+ "Warning: boolFilter was deprecated in v0.16.
+ Please use filter instead."
+ #-}
+
+-- Version 0.17
+
+idIsFold = id-is-foldr
+{-# WARNING_ON_USAGE idIsFold
+"Warning: idIsFold was deprecated in v0.17.
+Please use id-is-foldr instead."
+#-}
+++IsFold = ++-is-foldr
+{-# WARNING_ON_USAGE ++IsFold
+"Warning: ++IsFold was deprecated in v0.17.
+Please use ++-is-foldr instead."
+#-}
+mapIsFold = map-is-foldr
+{-# WARNING_ON_USAGE mapIsFold
+"Warning: mapIsFold was deprecated in v0.17.
+Please use map-is-foldr instead."
+#-}
diff --git a/src/Data/List/Relation/BagAndSetEquality.agda b/src/Data/List/Relation/BagAndSetEquality.agda
new file mode 100644
index 0000000..325671d
--- /dev/null
+++ b/src/Data/List/Relation/BagAndSetEquality.agda
@@ -0,0 +1,325 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Bag and set equality
+------------------------------------------------------------------------
+
+module Data.List.Relation.BagAndSetEquality where
+
+open import Algebra using (CommutativeSemiring; CommutativeMonoid)
+open import Algebra.FunctionProperties using (Idempotent)
+open import Category.Monad using (RawMonad)
+open import Data.List
+open import Data.List.Categorical using (monad; module MonadProperties)
+import Data.List.Properties as LP
+open import Data.List.Any using (Any; here; there)
+open import Data.List.Any.Properties
+open import Data.List.Membership.Propositional using (_∈_)
+open import Data.List.Relation.Subset.Propositional.Properties
+ using (⊆-preorder)
+open import Data.Product hiding (map)
+open import Data.Sum hiding (map)
+open import Data.Sum.Relation.Pointwise using (_⊎-cong_)
+open import Function
+open import Function.Equality using (_⟨$⟩_)
+import Function.Equivalence as FE
+open import Function.Inverse as Inv using (_↔_; Inverse; inverse)
+open import Function.Related as Related using (↔⇒; ⌊_⌋; ⌊_⌋→; ⇒→; SK-sym)
+open import Function.Related.TypeIsomorphisms
+open import Relation.Binary
+import Relation.Binary.EqReasoning as EqR
+import Relation.Binary.PreorderReasoning as PreorderReasoning
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≗_; refl)
+open import Relation.Nullary
+open import Data.List.Membership.Propositional.Properties
+
+------------------------------------------------------------------------
+-- Definitions
+
+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
+
+private
+ module Eq {k a} {A : Set a} = Setoid ([ k ]-Equality A)
+ module Ord {k a} {A : Set a} = Preorder ([ k ]-Order A)
+ open module ListMonad {ℓ} = RawMonad (monad {ℓ = ℓ})
+ module MP = MonadProperties
+
+------------------------------------------------------------------------
+-- 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 _⊆_ along with an additional relatedness
+
+module ⊆-Reasoning where
+ private
+ module PreOrder {a} {A : Set a} = PreorderReasoning (⊆-preorder A)
+
+ open PreOrder public
+ hiding (_≈⟨_⟩_) renaming (_∼⟨_⟩_ 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
+
+------------------------------------------------------------------------
+-- Congruence lemmas
+------------------------------------------------------------------------
+-- _∷_
+
+module _ {a k} {A : Set a} {x y : A} {xs ys} where
+
+ ∷-cong : x ≡ y → xs ∼[ k ] ys → x ∷ xs ∼[ k ] y ∷ ys
+ ∷-cong refl xs≈ys {y} =
+ y ∈ x ∷ xs ↔⟨ SK-sym $ ∷↔ (y ≡_) ⟩
+ (y ≡ x ⊎ y ∈ xs) ∼⟨ (y ≡ x ∎) ⊎-cong xs≈ys ⟩
+ (y ≡ x ⊎ y ∈ ys) ↔⟨ ∷↔ (y ≡_) ⟩
+ y ∈ x ∷ ys ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- map
+
+module _ {ℓ k} {A B : Set ℓ} {f g : A → B} {xs ys} where
+
+ map-cong : f ≗ g → xs ∼[ k ] ys → map f xs ∼[ k ] map g ys
+ map-cong f≗g xs≈ys {x} =
+ x ∈ map f xs ↔⟨ SK-sym $ map↔ ⟩
+ Any (λ y → x ≡ f y) xs ∼⟨ Any-cong (↔⇒ ∘ helper) xs≈ys ⟩
+ Any (λ y → x ≡ g y) ys ↔⟨ map↔ ⟩
+ x ∈ map g ys ∎
+ where
+ open Related.EquationalReasoning
+
+ helper : ∀ y → x ≡ f y ↔ x ≡ g y
+ helper y = record
+ { to = P.→-to-⟶ (λ x≡fy → P.trans x≡fy ( f≗g y))
+ ; from = P.→-to-⟶ (λ x≡gy → P.trans x≡gy (P.sym $ f≗g y))
+ ; inverse-of = record
+ { left-inverse-of = λ _ → P.≡-irrelevance _ _
+ ; right-inverse-of = λ _ → P.≡-irrelevance _ _
+ }
+ }
+
+------------------------------------------------------------------------
+-- _++_
+
+module _ {a k} {A : Set a} {xs₁ xs₂ ys₁ ys₂ : List A} where
+
+ ++-cong : xs₁ ∼[ k ] xs₂ → ys₁ ∼[ k ] ys₂ →
+ xs₁ ++ ys₁ ∼[ k ] xs₂ ++ ys₂
+ ++-cong xs₁≈xs₂ ys₁≈ys₂ {x} =
+ x ∈ xs₁ ++ ys₁ ↔⟨ SK-sym $ ++↔ ⟩
+ (x ∈ xs₁ ⊎ x ∈ ys₁) ∼⟨ xs₁≈xs₂ ⊎-cong ys₁≈ys₂ ⟩
+ (x ∈ xs₂ ⊎ x ∈ ys₂) ↔⟨ ++↔ ⟩
+ x ∈ xs₂ ++ ys₂ ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- concat
+
+module _ {a k} {A : Set a} {xss yss : List (List A)} where
+
+ concat-cong : xss ∼[ k ] yss → concat xss ∼[ k ] concat yss
+ concat-cong xss≈yss {x} =
+ x ∈ concat xss ↔⟨ SK-sym concat↔ ⟩
+ Any (Any (x ≡_)) xss ∼⟨ Any-cong (λ _ → _ ∎) xss≈yss ⟩
+ Any (Any (x ≡_)) yss ↔⟨ concat↔ ⟩
+ x ∈ concat yss ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- _>>=_
+
+module _ {ℓ k} {A B : Set ℓ} {xs ys} {f g : A → List B} where
+
+ >>=-cong : xs ∼[ k ] ys → (∀ x → f x ∼[ k ] g x) →
+ (xs >>= f) ∼[ k ] (ys >>= g)
+ >>=-cong xs≈ys f≈g {x} =
+ x ∈ (xs >>= f) ↔⟨ SK-sym >>=↔ ⟩
+ Any (λ y → x ∈ f y) xs ∼⟨ Any-cong (λ x → f≈g x) xs≈ys ⟩
+ Any (λ y → x ∈ g y) ys ↔⟨ >>=↔ ⟩
+ x ∈ (ys >>= g) ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- _⊛_
+
+module _ {ℓ k} {A B : Set ℓ} {fs gs : List (A → B)} {xs ys} where
+
+ ⊛-cong : fs ∼[ k ] gs → xs ∼[ k ] ys → (fs ⊛ xs) ∼[ k ] (gs ⊛ ys)
+ ⊛-cong fs≈gs xs≈ys =
+ >>=-cong fs≈gs λ f →
+ >>=-cong xs≈ys λ x →
+ _ ∎
+ where open Related.EquationalReasoning
+
+------------------------------------------------------------------------
+-- _⊗_
+
+module _ {ℓ k} {A B : Set ℓ} {xs₁ xs₂ : List A} {ys₁ ys₂ : List B} where
+
+ ⊗-cong : xs₁ ∼[ k ] xs₂ → ys₁ ∼[ k ] ys₂ →
+ (xs₁ ⊗ ys₁) ∼[ k ] (xs₂ ⊗ ys₂)
+ ⊗-cong xs₁≈xs₂ ys₁≈ys₂ =
+ ⊛-cong (⊛-cong (Ord.refl {x = [ _,_ ]}) xs₁≈xs₂) ys₁≈ys₂
+
+------------------------------------------------------------------------
+-- Other properties
+
+-- _++_ and [] form a commutative monoid, with either bag or set
+-- equality as the underlying equality.
+
+commutativeMonoid : ∀ {a} → Symmetric-kind → Set a →
+ CommutativeMonoid _ _
+commutativeMonoid {a} k A = record
+ { Carrier = List A
+ ; _≈_ = _∼[ ⌊ k ⌋ ]_
+ ; _∙_ = _++_
+ ; ε = []
+ ; isCommutativeMonoid = record
+ { isSemigroup = record
+ { isEquivalence = Eq.isEquivalence
+ ; assoc = λ xs ys zs →
+ Eq.reflexive (LP.++-assoc xs ys zs)
+ ; ∙-cong = ++-cong
+ }
+ ; identityˡ = λ xs {x} → x ∈ xs ∎
+ ; comm = λ xs ys {x} →
+ x ∈ xs ++ ys ↔⟨ ++↔++ xs ys ⟩
+ x ∈ ys ++ xs ∎
+ }
+ }
+ where open Related.EquationalReasoning
+
+-- The only list which is bag or set equal to the empty list (or a
+-- subset or subbag of the list) is the empty list itself.
+
+empty-unique : ∀ {k a} {A : Set a} {xs : List A} →
+ xs ∼[ ⌊ k ⌋→ ] [] → xs ≡ []
+empty-unique {xs = []} _ = refl
+empty-unique {xs = _ ∷ _} ∷∼[] with ⇒→ ∷∼[] (here refl)
+... | ()
+
+-- _++_ is idempotent (under set equality).
+
+++-idempotent : ∀ {a} {A : Set a} → Idempotent {A = List A} _∼[ set ]_ _++_
+++-idempotent {a} xs {x} =
+ x ∈ xs ++ xs ∼⟨ FE.equivalence ([ id , id ]′ ∘ _⟨$⟩_ (Inverse.from $ ++↔))
+ (_⟨$⟩_ (Inverse.to $ ++↔) ∘ inj₁) ⟩
+ x ∈ xs ∎
+ where open Related.EquationalReasoning
+
+-- The list monad's bind distributes from the left over _++_.
+
+>>=-left-distributive :
+ ∀ {ℓ} {A B : Set ℓ} (xs : List A) {f g : A → List B} →
+ (xs >>= λ x → f x ++ g x) ∼[ bag ] (xs >>= f) ++ (xs >>= g)
+>>=-left-distributive {ℓ} xs {f} {g} {y} =
+ y ∈ (xs >>= λ x → f x ++ g x) ↔⟨ SK-sym $ >>=↔ ⟩
+ Any (λ x → y ∈ f x ++ g x) xs ↔⟨ SK-sym (Any-cong (λ _ → ++↔) (_ ∎)) ⟩
+ Any (λ x → y ∈ f x ⊎ y ∈ g x) xs ↔⟨ SK-sym $ ⊎↔ ⟩
+ (Any (λ x → y ∈ f x) xs ⊎ Any (λ x → y ∈ g x) xs) ↔⟨ >>=↔ ⟨ _⊎-cong_ ⟩ >>=↔ ⟩
+ (y ∈ (xs >>= f) ⊎ y ∈ (xs >>= g)) ↔⟨ ++↔ ⟩
+ y ∈ (xs >>= f) ++ (xs >>= g) ∎
+ where open Related.EquationalReasoning
+
+-- The same applies to _⊛_.
+
+⊛-left-distributive :
+ ∀ {ℓ} {A B : Set ℓ} (fs : List (A → B)) xs₁ xs₂ →
+ (fs ⊛ (xs₁ ++ xs₂)) ∼[ bag ] (fs ⊛ xs₁) ++ (fs ⊛ xs₂)
+⊛-left-distributive {B = B} fs xs₁ xs₂ = begin
+ fs ⊛ (xs₁ ++ xs₂) ≡⟨⟩
+ (fs >>= λ f → xs₁ ++ xs₂ >>= return ∘ f) ≡⟨ (MP.cong (refl {x = fs}) λ f →
+ MP.right-distributive xs₁ xs₂ (return ∘ f)) ⟩
+ (fs >>= λ f → (xs₁ >>= return ∘ f) ++
+ (xs₂ >>= return ∘ f)) ≈⟨ >>=-left-distributive fs ⟩
+
+ (fs >>= λ f → xs₁ >>= return ∘ f) ++
+ (fs >>= λ f → xs₂ >>= return ∘ f) ≡⟨⟩
+
+ (fs ⊛ xs₁) ++ (fs ⊛ xs₂) ∎
+ where open EqR ([ bag ]-Equality B)
+
+private
+
+ -- If x ∷ xs is set equal to x ∷ ys, then xs and ys are not
+ -- necessarily set equal.
+
+ ¬-drop-cons : ∀ {a} {A : Set a} {x : A} →
+ ¬ (∀ {xs ys} → x ∷ xs ∼[ set ] x ∷ ys → xs ∼[ set ] ys)
+ ¬-drop-cons {x = x} drop-cons
+ with FE.Equivalence.to x∼[] ⟨$⟩ here refl
+ where
+ x,x≈x : (x ∷ x ∷ []) ∼[ set ] [ x ]
+ x,x≈x = ++-idempotent [ x ]
+
+ x∼[] : [ x ] ∼[ set ] []
+ x∼[] = drop-cons x,x≈x
+ ... | ()
+
+-- However, the corresponding property does hold for bag equality.
+
+drop-cons : ∀ {a} {A : Set a} {x : A} {xs ys} →
+ x ∷ xs ∼[ bag ] x ∷ ys → xs ∼[ bag ] ys
+drop-cons {x = x} eq = inverse (f eq) (f $ Inv.sym eq) (f∘f eq) (f∘f $ Inv.sym eq)
+ where
+ open Inverse
+ open P.≡-Reasoning
+
+ f : ∀ {xs ys z} → (z ∈ x ∷ xs) ↔ (z ∈ x ∷ ys) → z ∈ xs → z ∈ ys
+ f inv z∈xs with to inv ⟨$⟩ there z∈xs | left-inverse-of inv (there z∈xs)
+ ... | there z∈ys | left⁺ = z∈ys
+ ... | here refl | left⁺ with to inv ⟨$⟩ here refl | left-inverse-of inv (here refl)
+ ... | there z∈ys | left⁰ = z∈ys
+ ... | here refl | left⁰ with begin
+ here refl ≡⟨ P.sym left⁰ ⟩
+ from inv ⟨$⟩ here refl ≡⟨ left⁺ ⟩
+ there z∈xs ∎
+ ... | ()
+
+ f∘f : ∀ {xs ys z} (inv : (z ∈ x ∷ xs) ↔ (z ∈ x ∷ ys)) (p : z ∈ xs) →
+ f (Inv.sym inv) (f inv p) ≡ p
+ f∘f inv z∈xs with to inv ⟨$⟩ there z∈xs | left-inverse-of inv (there z∈xs)
+ f∘f inv z∈xs | there z∈ys | left⁺ with from inv ⟨$⟩ there z∈ys | right-inverse-of inv (there z∈ys)
+ f∘f inv z∈xs | there z∈ys | refl | .(there z∈xs) | _ = refl
+ f∘f inv z∈xs | here refl | left⁺ with to inv ⟨$⟩ here refl | left-inverse-of inv (here refl)
+ f∘f inv z∈xs | here refl | left⁺ | there z∈ys | left⁰ with from inv ⟨$⟩ there z∈ys | right-inverse-of inv (there z∈ys)
+ f∘f inv z∈xs | here refl | left⁺ | there z∈ys | refl | .(here refl) | _ with from inv ⟨$⟩ here refl
+ | right-inverse-of inv (here refl)
+ f∘f inv z∈xs | here refl | refl | there z∈ys | refl | .(here refl) | _ | .(there z∈xs) | _ = refl
+ f∘f inv z∈xs | here refl | left⁺ | here refl | left⁰ with begin
+ here refl ≡⟨ P.sym left⁰ ⟩
+ from inv ⟨$⟩ here refl ≡⟨ left⁺ ⟩
+ there z∈xs ∎
+ ... | ()
diff --git a/src/Data/List/Relation/Equality/DecPropositional.agda b/src/Data/List/Relation/Equality/DecPropositional.agda
new file mode 100644
index 0000000..695b7e1
--- /dev/null
+++ b/src/Data/List/Relation/Equality/DecPropositional.agda
@@ -0,0 +1,22 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Decidable equality over lists using propositional equality
+------------------------------------------------------------------------
+
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality
+
+module Data.List.Relation.Equality.DecPropositional
+ {a} {A : Set a} (_≟_ : Decidable {A = A} _≡_) where
+
+import Data.List.Relation.Equality.Propositional as PropositionalEq
+import Data.List.Relation.Equality.DecSetoid as DecSetoidEq
+
+------------------------------------------------------------------------
+-- Publically re-export everything from decSetoid and propositional
+-- equality
+
+open PropositionalEq public
+open DecSetoidEq (decSetoid _≟_) public
+ using (_≋?_; ≋-isDecEquivalence; ≋-decSetoid)
diff --git a/src/Data/List/Relation/Equality/DecSetoid.agda b/src/Data/List/Relation/Equality/DecSetoid.agda
new file mode 100644
index 0000000..ad577c1
--- /dev/null
+++ b/src/Data/List/Relation/Equality/DecSetoid.agda
@@ -0,0 +1,34 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Decidable equality over lists parameterised by some setoid
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Data.List.Relation.Equality.DecSetoid
+ {a ℓ} (DS : DecSetoid a ℓ) where
+
+import Data.List.Relation.Equality.Setoid as SetoidEquality
+import Data.List.Relation.Pointwise as PW
+open import Relation.Binary using (Decidable)
+open DecSetoid DS
+
+------------------------------------------------------------------------
+-- Make all definitions from setoid equality available
+
+open SetoidEquality setoid public
+
+------------------------------------------------------------------------
+-- Additional properties
+
+infix 4 _≋?_
+
+_≋?_ : Decidable _≋_
+_≋?_ = PW.decidable _≟_
+
+≋-isDecEquivalence : IsDecEquivalence _≋_
+≋-isDecEquivalence = PW.isDecEquivalence isDecEquivalence
+
+≋-decSetoid : DecSetoid a ℓ
+≋-decSetoid = PW.decSetoid DS
diff --git a/src/Data/List/Relation/Equality/Propositional.agda b/src/Data/List/Relation/Equality/Propositional.agda
new file mode 100644
index 0000000..becd445
--- /dev/null
+++ b/src/Data/List/Relation/Equality/Propositional.agda
@@ -0,0 +1,28 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Equality over lists using propositional equality
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Data.List.Relation.Equality.Propositional {a} {A : Set a} where
+
+open import Data.List
+import Data.List.Relation.Equality.Setoid as SetoidEquality
+open import Relation.Binary.PropositionalEquality
+
+------------------------------------------------------------------------
+-- Publically re-export everything from setoid equality
+
+open SetoidEquality (setoid A) public
+
+------------------------------------------------------------------------
+-- ≋ is propositional
+
+≋⇒≡ : _≋_ ⇒ _≡_
+≋⇒≡ [] = refl
+≋⇒≡ (refl ∷ xs≈ys) = cong (_ ∷_) (≋⇒≡ xs≈ys)
+
+≡⇒≋ : _≡_ ⇒ _≋_
+≡⇒≋ refl = ≋-refl
diff --git a/src/Data/List/Relation/Equality/Setoid.agda b/src/Data/List/Relation/Equality/Setoid.agda
new file mode 100644
index 0000000..f253203
--- /dev/null
+++ b/src/Data/List/Relation/Equality/Setoid.agda
@@ -0,0 +1,58 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Equality over lists parameterised by some setoid
+------------------------------------------------------------------------
+
+open import Relation.Binary using (Setoid)
+
+module Data.List.Relation.Equality.Setoid {a ℓ} (S : Setoid a ℓ) where
+
+open import Data.List.Base using (List)
+open import Level
+open import Relation.Binary renaming (Rel to Rel₂)
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Data.List.Relation.Pointwise as PW using (Pointwise)
+
+open Setoid S renaming (Carrier to A)
+
+------------------------------------------------------------------------
+-- Definition of equality
+
+infix 4 _≋_
+
+_≋_ : Rel₂ (List A) ℓ
+_≋_ = Pointwise _≈_
+
+open Pointwise public using ([]; _∷_)
+
+------------------------------------------------------------------------
+-- Relational properties
+
+≋-refl : Reflexive _≋_
+≋-refl = PW.refl refl
+
+≋-reflexive : _≡_ ⇒ _≋_
+≋-reflexive P.refl = ≋-refl
+
+≋-sym : Symmetric _≋_
+≋-sym = PW.symmetric sym
+
+≋-trans : Transitive _≋_
+≋-trans = PW.transitive trans
+
+≋-isEquivalence : IsEquivalence _≋_
+≋-isEquivalence = PW.isEquivalence isEquivalence
+
+≋-setoid : Setoid _ _
+≋-setoid = PW.setoid S
+
+------------------------------------------------------------------------
+-- Operations
+
+open PW public using
+ ( tabulate⁺
+ ; tabulate⁻
+ ; ++⁺
+ ; concat⁺
+ )
diff --git a/src/Data/List/Relation/Lex/Core.agda b/src/Data/List/Relation/Lex/Core.agda
new file mode 100644
index 0000000..7be15ba
--- /dev/null
+++ b/src/Data/List/Relation/Lex/Core.agda
@@ -0,0 +1,106 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Lexicographic ordering of lists
+------------------------------------------------------------------------
+
+module Data.List.Relation.Lex.Core where
+
+open import Data.Empty using (⊥; ⊥-elim)
+open import Data.Unit.Base using (⊤; tt)
+open import Function using (_∘_; flip; id)
+open import Data.Product using (_,_; proj₁; proj₂)
+open import Data.List.Base using (List; []; _∷_)
+open import Level using (_⊔_)
+open import Relation.Nullary using (Dec; yes; no; ¬_)
+open import Relation.Binary
+open import Data.List.Relation.Pointwise
+ using (Pointwise; []; _∷_; head; tail)
+
+-- The lexicographic ordering itself can be either strict or non-strict,
+-- depending on whether type P is inhabited.
+
+data Lex {a ℓ₁ ℓ₂} {A : Set a} (P : Set)
+ (_≈_ : Rel A ℓ₁) (_≺_ : Rel A ℓ₂) : Rel (List A) (ℓ₁ ⊔ ℓ₂) where
+ base : P → Lex P _≈_ _≺_ [] []
+ halt : ∀ {y ys} → Lex P _≈_ _≺_ [] (y ∷ ys)
+ this : ∀ {x xs y ys} (x≺y : x ≺ y) → Lex P _≈_ _≺_ (x ∷ xs) (y ∷ ys)
+ next : ∀ {x xs y ys} (x≈y : x ≈ y)
+ (xs<ys : Lex P _≈_ _≺_ xs ys) → Lex P _≈_ _≺_ (x ∷ xs) (y ∷ ys)
+
+-- Properties
+
+module _ {a ℓ₁ ℓ₂} {A : Set a} {P : Set}
+ {_≈_ : Rel A ℓ₁} {_≺_ : Rel A ℓ₂} where
+
+ private
+ _≋_ = Pointwise _≈_
+ _<_ = Lex P _≈_ _≺_
+
+ ¬≤-this : ∀ {x y xs ys} → ¬ (x ≈ y) → ¬ (x ≺ y) →
+ ¬ (x ∷ xs) < (y ∷ ys)
+ ¬≤-this x≉y x≮y (this x≺y) = x≮y x≺y
+ ¬≤-this x≉y x≮y (next x≈y xs<ys) = x≉y x≈y
+
+ ¬≤-next : ∀ {x y xs ys} → ¬ x ≺ y → ¬ xs < ys →
+ ¬ (x ∷ xs) < (y ∷ ys)
+ ¬≤-next x≮y xs≮ys (this x≺y) = x≮y x≺y
+ ¬≤-next x≮y xs≮ys (next _ xs<ys) = xs≮ys xs<ys
+
+ transitive : IsEquivalence _≈_ → _≺_ Respects₂ _≈_ → Transitive _≺_ →
+ Transitive _<_
+ transitive eq resp tr = trans
+ where
+ trans : Transitive (Lex P _≈_ _≺_)
+ trans (base p) (base _) = base p
+ trans (base y) halt = halt
+ trans halt (this y≺z) = halt
+ trans halt (next y≈z ys<zs) = halt
+ trans (this x≺y) (this y≺z) = this (tr x≺y y≺z)
+ trans (this x≺y) (next y≈z ys<zs) = this (proj₁ resp y≈z x≺y)
+ trans (next x≈y xs<ys) (this y≺z) =
+ this (proj₂ resp (IsEquivalence.sym eq x≈y) y≺z)
+ trans (next x≈y xs<ys) (next y≈z ys<zs) =
+ next (IsEquivalence.trans eq x≈y y≈z) (trans xs<ys ys<zs)
+
+ antisymmetric : Symmetric _≈_ → Irreflexive _≈_ _≺_ →
+ Asymmetric _≺_ → Antisymmetric _≋_ _<_
+ antisymmetric sym ir asym = as
+ where
+ as : Antisymmetric _≋_ _<_
+ as (base _) (base _) = []
+ as halt ()
+ as (this x≺y) (this y≺x) = ⊥-elim (asym x≺y y≺x)
+ as (this x≺y) (next y≈x ys<xs) = ⊥-elim (ir (sym y≈x) x≺y)
+ as (next x≈y xs<ys) (this y≺x) = ⊥-elim (ir (sym x≈y) y≺x)
+ as (next x≈y xs<ys) (next y≈x ys<xs) = x≈y ∷ as xs<ys ys<xs
+
+ respects₂ : IsEquivalence _≈_ → _≺_ Respects₂ _≈_ → _<_ Respects₂ _≋_
+ respects₂ eq (resp₁ , resp₂) = resp¹ , resp²
+ where
+ open IsEquivalence eq using (sym; trans)
+ resp¹ : ∀ {xs} → Lex P _≈_ _≺_ xs Respects _≋_
+ resp¹ [] xs<[] = xs<[]
+ resp¹ (_ ∷ _) halt = halt
+ resp¹ (x≈y ∷ _) (this z≺x) = this (resp₁ x≈y z≺x)
+ resp¹ (x≈y ∷ xs≋ys) (next z≈x zs<xs) =
+ next (trans z≈x x≈y) (resp¹ xs≋ys zs<xs)
+
+ resp² : ∀ {ys} → flip (Lex P _≈_ _≺_) ys Respects _≋_
+ resp² [] []<ys = []<ys
+ resp² (x≈z ∷ _) (this x≺y) = this (resp₂ x≈z x≺y)
+ resp² (x≈z ∷ xs≋zs) (next x≈y xs<ys) =
+ next (trans (sym x≈z) x≈y) (resp² xs≋zs xs<ys)
+
+ decidable : Dec P → Decidable _≈_ → Decidable _≺_ → Decidable _<_
+ decidable (yes p) dec-≈ dec-≺ [] [] = yes (base p)
+ decidable (no ¬p) dec-≈ dec-≺ [] [] = no λ{(base p) → ¬p p}
+ decidable dec-P dec-≈ dec-≺ [] (y ∷ ys) = yes halt
+ decidable dec-P dec-≈ dec-≺ (x ∷ xs) [] = no λ()
+ decidable dec-P dec-≈ dec-≺ (x ∷ xs) (y ∷ ys) with dec-≺ x y
+ ... | yes x≺y = yes (this x≺y)
+ ... | no x≮y with dec-≈ x y
+ ... | no x≉y = no (¬≤-this x≉y x≮y)
+ ... | yes x≈y with decidable dec-P dec-≈ dec-≺ xs ys
+ ... | yes xs<ys = yes (next x≈y xs<ys)
+ ... | no xs≮ys = no (¬≤-next x≮y xs≮ys)
diff --git a/src/Data/List/Relation/Lex/NonStrict.agda b/src/Data/List/Relation/Lex/NonStrict.agda
new file mode 100644
index 0000000..648a8e9
--- /dev/null
+++ b/src/Data/List/Relation/Lex/NonStrict.agda
@@ -0,0 +1,188 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Lexicographic ordering of lists
+------------------------------------------------------------------------
+
+-- The definitions of lexicographic orderings used here is suitable if
+-- the argument order is a (non-strict) partial order.
+
+module Data.List.Relation.Lex.NonStrict where
+
+open import Data.Empty using (⊥)
+open import Function
+open import Data.Unit.Base using (⊤; tt)
+open import Data.Product
+open import Data.List.Base
+open import Data.List.Relation.Pointwise using (Pointwise; [])
+import Data.List.Relation.Lex.Strict as Strict
+open import Level
+open import Relation.Nullary
+open import Relation.Binary
+import Relation.Binary.Construct.NonStrictToStrict as Conv
+
+------------------------------------------------------------------------
+-- Publically re-export definitions from Core
+
+open import Data.List.Relation.Lex.Core as Core public
+ using (base; halt; this; next; ²-this; ²-next)
+
+------------------------------------------------------------------------
+-- Strict lexicographic ordering.
+
+module _ {a ℓ₁ ℓ₂} {A : Set a} where
+
+ Lex-< : (_≈_ : Rel A ℓ₁) (_≼_ : Rel A ℓ₂) → Rel (List A) (ℓ₁ ⊔ ℓ₂)
+ Lex-< _≈_ _≼_ = Core.Lex ⊥ _≈_ (Conv._<_ _≈_ _≼_)
+
+ -- Properties
+
+ module _ {_≈_ : Rel A ℓ₁} {_≼_ : Rel A ℓ₂} where
+
+ private
+ _≋_ = Pointwise _≈_
+ _<_ = Lex-< _≈_ _≼_
+
+ <-irreflexive : Irreflexive _≋_ _<_
+ <-irreflexive = Strict.<-irreflexive (Conv.<-irrefl _≈_ _≼_)
+
+ <-asymmetric : IsEquivalence _≈_ → _≼_ Respects₂ _≈_ →
+ Antisymmetric _≈_ _≼_ → Asymmetric _<_
+ <-asymmetric eq resp antisym =
+ Strict.<-asymmetric sym (Conv.<-resp-≈ _ _ eq resp)
+ (Conv.<-asym _≈_ _ antisym)
+ where open IsEquivalence eq
+
+ <-antisymmetric : Symmetric _≈_ → Antisymmetric _≈_ _≼_ →
+ Antisymmetric _≋_ _<_
+ <-antisymmetric sym antisym =
+ Core.antisymmetric sym
+ (Conv.<-irrefl _≈_ _≼_)
+ (Conv.<-asym _ _≼_ antisym)
+
+ <-transitive : IsPartialOrder _≈_ _≼_ → Transitive _<_
+ <-transitive po =
+ Core.transitive isEquivalence
+ (Conv.<-resp-≈ _ _ isEquivalence ≤-resp-≈)
+ (Conv.<-trans _ _≼_ po)
+ where open IsPartialOrder po
+
+ <-resp₂ : IsEquivalence _≈_ → _≼_ Respects₂ _≈_ → _<_ Respects₂ _≋_
+ <-resp₂ eq resp = Core.respects₂ eq (Conv.<-resp-≈ _ _ eq resp)
+
+ <-compare : Symmetric _≈_ → Decidable _≈_ → Antisymmetric _≈_ _≼_ →
+ Total _≼_ → Trichotomous _≋_ _<_
+ <-compare sym _≟_ antisym tot =
+ Strict.<-compare sym (Conv.<-trichotomous _ _ sym _≟_ antisym tot)
+
+ <-decidable : Decidable _≈_ → Decidable _≼_ → Decidable _<_
+ <-decidable _≟_ _≼?_ =
+ Core.decidable (no id) _≟_ (Conv.<-decidable _ _ _≟_ _≼?_)
+
+ <-isStrictPartialOrder : IsPartialOrder _≈_ _≼_ →
+ IsStrictPartialOrder _≋_ _<_
+ <-isStrictPartialOrder po =
+ Strict.<-isStrictPartialOrder
+ (Conv.<-isStrictPartialOrder _ _ po)
+
+ <-isStrictTotalOrder : Decidable _≈_ → IsTotalOrder _≈_ _≼_ →
+ IsStrictTotalOrder _≋_ _<_
+ <-isStrictTotalOrder dec tot =
+ Strict.<-isStrictTotalOrder
+ (Conv.<-isStrictTotalOrder₁ _ _ dec tot)
+
+<-strictPartialOrder : ∀ {a ℓ₁ ℓ₂} → Poset a ℓ₁ ℓ₂ →
+ StrictPartialOrder _ _ _
+<-strictPartialOrder po = record
+ { isStrictPartialOrder = <-isStrictPartialOrder isPartialOrder
+ } where open Poset po
+
+<-strictTotalOrder : ∀ {a ℓ₁ ℓ₂} → DecTotalOrder a ℓ₁ ℓ₂ →
+ StrictTotalOrder _ _ _
+<-strictTotalOrder dtot = record
+ { isStrictTotalOrder = <-isStrictTotalOrder _≟_ isTotalOrder
+ } where open DecTotalOrder dtot
+
+------------------------------------------------------------------------
+-- Non-strict lexicographic ordering.
+
+module _ {a ℓ₁ ℓ₂} {A : Set a} where
+
+ Lex-≤ : (_≈_ : Rel A ℓ₁) (_≼_ : Rel A ℓ₂) → Rel (List A) (ℓ₁ ⊔ ℓ₂)
+ Lex-≤ _≈_ _≼_ = Core.Lex ⊤ _≈_ (Conv._<_ _≈_ _≼_)
+
+ ≤-reflexive : ∀ _≈_ _≼_ → Pointwise _≈_ ⇒ Lex-≤ _≈_ _≼_
+ ≤-reflexive _≈_ _≼_ = Strict.≤-reflexive _≈_ (Conv._<_ _≈_ _≼_)
+
+ -- Properties
+ module _ {_≈_ : Rel A ℓ₁} {_≼_ : Rel A ℓ₂} where
+
+ private
+ _≋_ = Pointwise _≈_
+ _≤_ = Lex-≤ _≈_ _≼_
+
+ ≤-antisymmetric : Symmetric _≈_ → Antisymmetric _≈_ _≼_ →
+ Antisymmetric _≋_ _≤_
+ ≤-antisymmetric sym antisym =
+ Core.antisymmetric sym
+ (Conv.<-irrefl _≈_ _≼_)
+ (Conv.<-asym _ _≼_ antisym)
+
+ ≤-transitive : IsPartialOrder _≈_ _≼_ → Transitive _≤_
+ ≤-transitive po =
+ Core.transitive isEquivalence
+ (Conv.<-resp-≈ _ _ isEquivalence ≤-resp-≈)
+ (Conv.<-trans _ _≼_ po)
+ where open IsPartialOrder po
+
+ ≤-resp₂ : IsEquivalence _≈_ → _≼_ Respects₂ _≈_ → _≤_ Respects₂ _≋_
+ ≤-resp₂ eq resp = Core.respects₂ eq (Conv.<-resp-≈ _ _ eq resp)
+
+ ≤-decidable : Decidable _≈_ → Decidable _≼_ → Decidable _≤_
+ ≤-decidable _≟_ _≼?_ =
+ Core.decidable (yes tt) _≟_ (Conv.<-decidable _ _ _≟_ _≼?_)
+
+ ≤-total : Symmetric _≈_ → Decidable _≈_ → Antisymmetric _≈_ _≼_ →
+ Total _≼_ → Total _≤_
+ ≤-total sym dec-≈ antisym tot =
+ Strict.≤-total sym (Conv.<-trichotomous _ _ sym dec-≈ antisym tot)
+
+ ≤-isPreorder : IsPartialOrder _≈_ _≼_ → IsPreorder _≋_ _≤_
+ ≤-isPreorder po =
+ Strict.≤-isPreorder
+ isEquivalence (Conv.<-trans _ _ po)
+ (Conv.<-resp-≈ _ _ isEquivalence ≤-resp-≈)
+ where open IsPartialOrder po
+
+ ≤-isPartialOrder : IsPartialOrder _≈_ _≼_ → IsPartialOrder _≋_ _≤_
+ ≤-isPartialOrder po =
+ Strict.≤-isPartialOrder
+ (Conv.<-isStrictPartialOrder _ _ po)
+
+ ≤-isTotalOrder : Decidable _≈_ → IsTotalOrder _≈_ _≼_ →
+ IsTotalOrder _≋_ _≤_
+ ≤-isTotalOrder dec tot =
+ Strict.≤-isTotalOrder
+ (Conv.<-isStrictTotalOrder₁ _ _ dec tot)
+
+ ≤-isDecTotalOrder : IsDecTotalOrder _≈_ _≼_ →
+ IsDecTotalOrder _≋_ _≤_
+ ≤-isDecTotalOrder dtot =
+ Strict.≤-isDecTotalOrder
+ (Conv.<-isStrictTotalOrder₂ _ _ dtot)
+
+≤-preorder : ∀ {a ℓ₁ ℓ₂} → Poset a ℓ₁ ℓ₂ → Preorder _ _ _
+≤-preorder po = record
+ { isPreorder = ≤-isPreorder isPartialOrder
+ } where open Poset po
+
+≤-partialOrder : ∀ {a ℓ₁ ℓ₂} → Poset a ℓ₁ ℓ₂ → Poset _ _ _
+≤-partialOrder po = record
+ { isPartialOrder = ≤-isPartialOrder isPartialOrder
+ } where open Poset po
+
+≤-decTotalOrder : ∀ {a ℓ₁ ℓ₂} → DecTotalOrder a ℓ₁ ℓ₂ →
+ DecTotalOrder _ _ _
+≤-decTotalOrder dtot = record
+ { isDecTotalOrder = ≤-isDecTotalOrder isDecTotalOrder
+ } where open DecTotalOrder dtot
diff --git a/src/Data/List/Relation/Lex/Strict.agda b/src/Data/List/Relation/Lex/Strict.agda
new file mode 100644
index 0000000..4c08022
--- /dev/null
+++ b/src/Data/List/Relation/Lex/Strict.agda
@@ -0,0 +1,227 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Lexicographic ordering of lists
+------------------------------------------------------------------------
+
+-- The definitions of lexicographic ordering used here are suitable if
+-- the argument order is a strict partial order.
+
+module Data.List.Relation.Lex.Strict where
+
+open import Data.Empty using (⊥)
+open import Data.Unit.Base using (⊤; tt)
+open import Function using (_∘_; id)
+open import Data.Product using (_,_)
+open import Data.Sum using (inj₁; inj₂)
+open import Data.List.Base using (List; []; _∷_)
+open import Level using (_⊔_)
+open import Relation.Nullary using (yes; no; ¬_)
+open import Relation.Binary
+open import Relation.Binary.Consequences
+open import Data.List.Relation.Pointwise as Pointwise
+ using (Pointwise; []; _∷_; head; tail)
+
+open import Data.List.Relation.Lex.Core as Core public
+ using (base; halt; this; next; ²-this; ²-next)
+
+----------------------------------------------------------------------
+-- Strict lexicographic ordering.
+
+module _ {a ℓ₁ ℓ₂} {A : Set a} where
+
+ Lex-< : (_≈_ : Rel A ℓ₁) (_≺_ : Rel A ℓ₂) → Rel (List A) (ℓ₁ ⊔ ℓ₂)
+ Lex-< = Core.Lex ⊥
+
+ -- Properties
+
+ module _ {_≈_ : Rel A ℓ₁} {_≺_ : Rel A ℓ₂} where
+
+ private
+ _≋_ = Pointwise _≈_
+ _<_ = Lex-< _≈_ _≺_
+
+ ¬[]<[] : ¬ [] < []
+ ¬[]<[] (base ())
+
+ <-irreflexive : Irreflexive _≈_ _≺_ → Irreflexive _≋_ _<_
+ <-irreflexive irr [] (base ())
+ <-irreflexive irr (x≈y ∷ xs≋ys) (this x<y) = irr x≈y x<y
+ <-irreflexive irr (x≈y ∷ xs≋ys) (next _ xs⊴ys) =
+ <-irreflexive irr xs≋ys xs⊴ys
+
+ <-asymmetric : Symmetric _≈_ → _≺_ Respects₂ _≈_ → Asymmetric _≺_ →
+ Asymmetric _<_
+ <-asymmetric sym resp as = asym
+ where
+ irrefl : Irreflexive _≈_ _≺_
+ irrefl = asym⟶irr resp sym as
+
+ asym : Asymmetric _<_
+ asym (halt) ()
+ asym (base bot) _ = bot
+ asym (this x<y) (this y<x) = as x<y y<x
+ asym (this x<y) (next y≈x ys⊴xs) = irrefl (sym y≈x) x<y
+ asym (next x≈y xs⊴ys) (this y<x) = irrefl (sym x≈y) y<x
+ asym (next x≈y xs⊴ys) (next y≈x ys⊴xs) = asym xs⊴ys ys⊴xs
+
+ <-antisymmetric : Symmetric _≈_ → Irreflexive _≈_ _≺_ →
+ Asymmetric _≺_ → Antisymmetric _≋_ _<_
+ <-antisymmetric = Core.antisymmetric
+
+ <-transitive : IsEquivalence _≈_ → _≺_ Respects₂ _≈_ →
+ Transitive _≺_ → Transitive _<_
+ <-transitive = Core.transitive
+
+ <-compare : Symmetric _≈_ → Trichotomous _≈_ _≺_ →
+ Trichotomous _≋_ _<_
+ <-compare sym tri [] [] = tri≈ ¬[]<[] [] ¬[]<[]
+ <-compare sym tri [] (y ∷ ys) = tri< halt (λ()) (λ())
+ <-compare sym tri (x ∷ xs) [] = tri> (λ()) (λ()) halt
+ <-compare sym tri (x ∷ xs) (y ∷ ys) with tri x y
+ ... | tri< x<y x≉y y≮x =
+ tri< (this x<y) (x≉y ∘ head) (¬≤-this (x≉y ∘ sym) y≮x)
+ ... | tri> x≮y x≉y y<x =
+ tri> (¬≤-this x≉y x≮y) (x≉y ∘ head) (this y<x)
+ ... | tri≈ x≮y x≈y y≮x with <-compare sym tri xs ys
+ ... | tri< xs<ys xs≉ys ys≮xs =
+ tri< (next x≈y xs<ys) (xs≉ys ∘ tail) (¬≤-next y≮x ys≮xs)
+ ... | tri≈ xs≮ys xs≈ys ys≮xs =
+ tri≈ (¬≤-next x≮y xs≮ys) (x≈y ∷ xs≈ys) (¬≤-next y≮x ys≮xs)
+ ... | tri> xs≮ys xs≉ys ys<xs =
+ tri> (¬≤-next x≮y xs≮ys) (xs≉ys ∘ tail) (next (sym x≈y) ys<xs)
+
+ <-decidable : Decidable _≈_ → Decidable _≺_ → Decidable _<_
+ <-decidable = Core.decidable (no id)
+
+ <-respects₂ : IsEquivalence _≈_ → _≺_ Respects₂ _≈_ →
+ _<_ Respects₂ _≋_
+ <-respects₂ = Core.respects₂
+
+ <-isStrictPartialOrder : IsStrictPartialOrder _≈_ _≺_ →
+ IsStrictPartialOrder _≋_ _<_
+ <-isStrictPartialOrder spo = record
+ { isEquivalence = Pointwise.isEquivalence isEquivalence
+ ; irrefl = <-irreflexive irrefl
+ ; trans = Core.transitive isEquivalence <-resp-≈ trans
+ ; <-resp-≈ = Core.respects₂ isEquivalence <-resp-≈
+ } where open IsStrictPartialOrder spo
+
+ <-isStrictTotalOrder : IsStrictTotalOrder _≈_ _≺_ →
+ IsStrictTotalOrder _≋_ _<_
+ <-isStrictTotalOrder sto = record
+ { isEquivalence = Pointwise.isEquivalence isEquivalence
+ ; trans = <-transitive isEquivalence <-resp-≈ trans
+ ; compare = <-compare Eq.sym compare
+ } where open IsStrictTotalOrder sto
+
+<-strictPartialOrder : ∀ {a ℓ₁ ℓ₂} → StrictPartialOrder a ℓ₁ ℓ₂ →
+ StrictPartialOrder _ _ _
+<-strictPartialOrder spo = record
+ { isStrictPartialOrder = <-isStrictPartialOrder isStrictPartialOrder
+ } where open StrictPartialOrder spo
+
+<-strictTotalOrder : ∀ {a ℓ₁ ℓ₂} → StrictTotalOrder a ℓ₁ ℓ₂ →
+ StrictTotalOrder _ _ _
+<-strictTotalOrder sto = record
+ { isStrictTotalOrder = <-isStrictTotalOrder isStrictTotalOrder
+ } where open StrictTotalOrder sto
+
+----------------------------------------------------------------------
+-- Non-strict lexicographic ordering.
+
+module _ {a ℓ₁ ℓ₂} {A : Set a} where
+
+ Lex-≤ : (_≈_ : Rel A ℓ₁) (_≺_ : Rel A ℓ₂) → Rel (List A) (ℓ₁ ⊔ ℓ₂)
+ Lex-≤ = Core.Lex ⊤
+
+ -- Properties
+
+ ≤-reflexive : (_≈_ : Rel A ℓ₁) (_≺_ : Rel A ℓ₂) →
+ Pointwise _≈_ ⇒ Lex-≤ _≈_ _≺_
+ ≤-reflexive _≈_ _≺_ [] = base tt
+ ≤-reflexive _≈_ _≺_ (x≈y ∷ xs≈ys) =
+ next x≈y (≤-reflexive _≈_ _≺_ xs≈ys)
+
+ module _ {_≈_ : Rel A ℓ₁} {_≺_ : Rel A ℓ₂} where
+
+ private
+ _≋_ = Pointwise _≈_
+ _≤_ = Lex-≤ _≈_ _≺_
+
+ ≤-antisymmetric : Symmetric _≈_ → Irreflexive _≈_ _≺_ →
+ Asymmetric _≺_ → Antisymmetric _≋_ _≤_
+ ≤-antisymmetric = Core.antisymmetric
+
+ ≤-transitive : IsEquivalence _≈_ → _≺_ Respects₂ _≈_ →
+ Transitive _≺_ → Transitive _≤_
+ ≤-transitive = Core.transitive
+
+ -- Note that trichotomy is an unnecessarily strong precondition for
+ -- the following lemma.
+
+ ≤-total : Symmetric _≈_ → Trichotomous _≈_ _≺_ → Total _≤_
+ ≤-total _ _ [] [] = inj₁ (base tt)
+ ≤-total _ _ [] (x ∷ xs) = inj₁ halt
+ ≤-total _ _ (x ∷ xs) [] = inj₂ halt
+ ≤-total sym tri (x ∷ xs) (y ∷ ys) with tri x y
+ ... | tri< x<y _ _ = inj₁ (this x<y)
+ ... | tri> _ _ y<x = inj₂ (this y<x)
+ ... | tri≈ _ x≈y _ with ≤-total sym tri xs ys
+ ... | inj₁ xs≲ys = inj₁ (next x≈y xs≲ys)
+ ... | inj₂ ys≲xs = inj₂ (next (sym x≈y) ys≲xs)
+
+ ≤-decidable : Decidable _≈_ → Decidable _≺_ → Decidable _≤_
+ ≤-decidable = Core.decidable (yes tt)
+
+ ≤-respects₂ : IsEquivalence _≈_ → _≺_ Respects₂ _≈_ →
+ _≤_ Respects₂ _≋_
+ ≤-respects₂ = Core.respects₂
+
+ ≤-isPreorder : IsEquivalence _≈_ → Transitive _≺_ →
+ _≺_ Respects₂ _≈_ → IsPreorder _≋_ _≤_
+ ≤-isPreorder eq tr resp = record
+ { isEquivalence = Pointwise.isEquivalence eq
+ ; reflexive = ≤-reflexive _≈_ _≺_
+ ; trans = Core.transitive eq resp tr
+ }
+
+ ≤-isPartialOrder : IsStrictPartialOrder _≈_ _≺_ →
+ IsPartialOrder _≋_ _≤_
+ ≤-isPartialOrder spo = record
+ { isPreorder = ≤-isPreorder isEquivalence trans <-resp-≈
+ ; antisym = Core.antisymmetric Eq.sym irrefl asym
+ }
+ where open IsStrictPartialOrder spo
+
+ ≤-isTotalOrder : IsStrictTotalOrder _≈_ _≺_ → IsTotalOrder _≋_ _≤_
+ ≤-isTotalOrder sto = record
+ { isPartialOrder = ≤-isPartialOrder isStrictPartialOrder
+ ; total = ≤-total Eq.sym compare
+ }
+ where open IsStrictTotalOrder sto
+
+ ≤-isDecTotalOrder : IsStrictTotalOrder _≈_ _≺_ →
+ IsDecTotalOrder _≋_ _≤_
+ ≤-isDecTotalOrder sto = record
+ { isTotalOrder = ≤-isTotalOrder sto
+ ; _≟_ = Pointwise.decidable _≟_
+ ; _≤?_ = ≤-decidable _≟_ _<?_
+ }
+ where open IsStrictTotalOrder sto
+
+≤-preorder : ∀ {a ℓ₁ ℓ₂} → Preorder a ℓ₁ ℓ₂ → Preorder _ _ _
+≤-preorder pre = record
+ { isPreorder = ≤-isPreorder isEquivalence trans ∼-resp-≈
+ } where open Preorder pre
+
+≤-partialOrder : ∀ {a ℓ₁ ℓ₂} → StrictPartialOrder a ℓ₁ ℓ₂ → Poset _ _ _
+≤-partialOrder spo = record
+ { isPartialOrder = ≤-isPartialOrder isStrictPartialOrder
+ } where open StrictPartialOrder spo
+
+≤-decTotalOrder : ∀ {a ℓ₁ ℓ₂} → StrictTotalOrder a ℓ₁ ℓ₂ →
+ DecTotalOrder _ _ _
+≤-decTotalOrder sto = record
+ { isDecTotalOrder = ≤-isDecTotalOrder isStrictTotalOrder
+ } where open StrictTotalOrder sto
diff --git a/src/Data/List/Relation/Permutation/Inductive.agda b/src/Data/List/Relation/Permutation/Inductive.agda
new file mode 100644
index 0000000..ffb6968
--- /dev/null
+++ b/src/Data/List/Relation/Permutation/Inductive.agda
@@ -0,0 +1,78 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- An inductive definition for the permutation relation
+------------------------------------------------------------------------
+
+module Data.List.Relation.Permutation.Inductive {a} {A : Set a} where
+
+open import Data.List using (List; []; _∷_)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality using (_≡_; refl)
+import Relation.Binary.EqReasoning as EqReasoning
+
+------------------------------------------------------------------------
+-- An inductive definition of permutation
+
+infix 3 _↭_
+
+data _↭_ : Rel (List A) a where
+ refl : ∀ {xs} → xs ↭ xs
+ prep : ∀ {xs ys} x → xs ↭ ys → x ∷ xs ↭ x ∷ ys
+ swap : ∀ {xs ys} x y → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys
+ trans : ∀ {xs ys zs} → xs ↭ ys → ys ↭ zs → xs ↭ zs
+
+------------------------------------------------------------------------
+-- _↭_ is an equivalence
+
+↭-reflexive : _≡_ ⇒ _↭_
+↭-reflexive refl = refl
+
+↭-refl : Reflexive _↭_
+↭-refl = refl
+
+↭-sym : ∀ {xs ys} → xs ↭ ys → ys ↭ xs
+↭-sym refl = refl
+↭-sym (prep x xs↭ys) = prep x (↭-sym xs↭ys)
+↭-sym (swap x y xs↭ys) = swap y x (↭-sym xs↭ys)
+↭-sym (trans xs↭ys ys↭zs) = trans (↭-sym ys↭zs) (↭-sym xs↭ys)
+
+↭-trans : Transitive _↭_
+↭-trans = trans
+
+↭-isEquivalence : IsEquivalence _↭_
+↭-isEquivalence = record
+ { refl = refl
+ ; sym = ↭-sym
+ ; trans = trans
+ }
+
+↭-setoid : Setoid _ _
+↭-setoid = record
+ { isEquivalence = ↭-isEquivalence
+ }
+
+------------------------------------------------------------------------
+-- A reasoning API to chain permutation proofs and allow "zooming in"
+-- to localised reasoning.
+
+module PermutationReasoning where
+
+ open EqReasoning ↭-setoid
+ using (_IsRelatedTo_; relTo)
+
+ open EqReasoning ↭-setoid public
+ using (begin_ ; _∎ ; _≡⟨⟩_; _≡⟨_⟩_)
+ renaming (_≈⟨_⟩_ to _↭⟨_⟩_)
+
+ infixr 2 _∷_<⟨_⟩_ _∷_∷_<<⟨_⟩_
+
+ -- Skip reasoning on the first element
+ _∷_<⟨_⟩_ : ∀ x xs {ys zs : List A} → xs ↭ ys →
+ (x ∷ ys) IsRelatedTo zs → (x ∷ xs) IsRelatedTo zs
+ x ∷ xs <⟨ xs↭ys ⟩ rel = relTo (trans (prep x xs↭ys) (begin rel))
+
+ -- Skip reasoning about the first two elements
+ _∷_∷_<<⟨_⟩_ : ∀ x y xs {ys zs : List A} → xs ↭ ys →
+ (y ∷ x ∷ ys) IsRelatedTo zs → (x ∷ y ∷ xs) IsRelatedTo zs
+ x ∷ y ∷ xs <<⟨ xs↭ys ⟩ rel = relTo (trans (swap x y xs↭ys) (begin rel))
diff --git a/src/Data/List/Relation/Permutation/Inductive/Properties.agda b/src/Data/List/Relation/Permutation/Inductive/Properties.agda
new file mode 100644
index 0000000..2966c72
--- /dev/null
+++ b/src/Data/List/Relation/Permutation/Inductive/Properties.agda
@@ -0,0 +1,274 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of permutation
+------------------------------------------------------------------------
+
+module Data.List.Relation.Permutation.Inductive.Properties where
+
+open import Algebra
+open import Algebra.FunctionProperties
+open import Algebra.Structures
+open import Data.List.Base as List
+open import Data.List.Relation.Permutation.Inductive
+open import Data.List.Any using (Any; here; there)
+open import Data.List.All using (All; []; _∷_)
+open import Data.List.Membership.Propositional
+open import Data.List.Membership.Propositional.Properties
+open import Data.List.Relation.BagAndSetEquality
+ using (bag; _∼[_]_; empty-unique; drop-cons; commutativeMonoid)
+import Data.List.Properties as Lₚ
+open import Data.Product using (_,_; _×_; ∃; ∃₂)
+open import Function using (_∘_)
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Inverse as Inv using (inverse)
+open import Relation.Unary using (Pred)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as ≡
+ using (_≡_ ; refl ; cong; cong₂; _≢_; inspect)
+
+open PermutationReasoning
+
+------------------------------------------------------------------------
+-- sym
+
+module _ {a} {A : Set a} where
+
+ ↭-sym-involutive : ∀ {xs ys : List A} (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p
+ ↭-sym-involutive refl = refl
+ ↭-sym-involutive (prep x ↭) = cong (prep x) (↭-sym-involutive ↭)
+ ↭-sym-involutive (swap x y ↭) = cong (swap x y) (↭-sym-involutive ↭)
+ ↭-sym-involutive (trans ↭₁ ↭₂) =
+ cong₂ trans (↭-sym-involutive ↭₁) (↭-sym-involutive ↭₂)
+
+------------------------------------------------------------------------
+-- Relationships to other predicates
+
+module _ {a} {A : Set a} where
+
+ All-resp-↭ : ∀ {ℓ} {P : Pred A ℓ} → (All P) Respects _↭_
+ All-resp-↭ refl wit = wit
+ All-resp-↭ (prep x p) (px ∷ wit) = px ∷ All-resp-↭ p wit
+ All-resp-↭ (swap x y p) (px ∷ py ∷ wit) = py ∷ px ∷ All-resp-↭ p wit
+ All-resp-↭ (trans p₁ p₂) wit = All-resp-↭ p₂ (All-resp-↭ p₁ wit)
+
+ Any-resp-↭ : ∀ {ℓ} {P : Pred A ℓ} → (Any P) Respects _↭_
+ Any-resp-↭ refl wit = wit
+ Any-resp-↭ (prep x p) (here px) = here px
+ Any-resp-↭ (prep x p) (there wit) = there (Any-resp-↭ p wit)
+ Any-resp-↭ (swap x y p) (here px) = there (here px)
+ Any-resp-↭ (swap x y p) (there (here px)) = here px
+ Any-resp-↭ (swap x y p) (there (there wit)) = there (there (Any-resp-↭ p wit))
+ Any-resp-↭ (trans p p₁) wit = Any-resp-↭ p₁ (Any-resp-↭ p wit)
+
+ ∈-resp-↭ : ∀ {x : A} → (x ∈_) Respects _↭_
+ ∈-resp-↭ = Any-resp-↭
+
+------------------------------------------------------------------------
+-- map
+
+module _ {a b} {A : Set a} {B : Set b} (f : A → B) where
+
+ map⁺ : ∀ {xs ys} → xs ↭ ys → map f xs ↭ map f ys
+ map⁺ refl = refl
+ map⁺ (prep x p) = prep _ (map⁺ p)
+ map⁺ (swap x y p) = swap _ _ (map⁺ p)
+ map⁺ (trans p₁ p₂) = trans (map⁺ p₁) (map⁺ p₂)
+
+------------------------------------------------------------------------
+-- _++_
+
+module _ {a} {A : Set a} where
+
+
+ ++⁺ˡ : ∀ xs {ys zs : List A} → ys ↭ zs → xs ++ ys ↭ xs ++ zs
+ ++⁺ˡ [] ys↭zs = ys↭zs
+ ++⁺ˡ (x ∷ xs) ys↭zs = prep x (++⁺ˡ xs ys↭zs)
+
+ ++⁺ʳ : ∀ {xs ys : List A} zs → xs ↭ ys → xs ++ zs ↭ ys ++ zs
+ ++⁺ʳ zs refl = refl
+ ++⁺ʳ zs (prep x ↭) = prep x (++⁺ʳ zs ↭)
+ ++⁺ʳ zs (swap x y ↭) = swap x y (++⁺ʳ zs ↭)
+ ++⁺ʳ zs (trans ↭₁ ↭₂) = trans (++⁺ʳ zs ↭₁) (++⁺ʳ zs ↭₂)
+
+ ++⁺ : _++_ Preserves₂ _↭_ ⟶ _↭_ ⟶ _↭_
+ ++⁺ ws↭xs ys↭zs = trans (++⁺ʳ _ ws↭xs) (++⁺ˡ _ ys↭zs)
+
+ -- Some useful lemmas
+
+ zoom : ∀ h {t xs ys : List A} → xs ↭ ys → h ++ xs ++ t ↭ h ++ ys ++ t
+ zoom h {t} = ++⁺ˡ h ∘ ++⁺ʳ t
+
+ inject : ∀ (v : A) {ws xs ys zs} → ws ↭ ys → xs ↭ zs →
+ ws ++ [ v ] ++ xs ↭ ys ++ [ v ] ++ zs
+ inject v ws↭ys xs↭zs = trans (++⁺ˡ _ (prep v xs↭zs)) (++⁺ʳ _ ws↭ys)
+
+ shift : ∀ v (xs ys : List A) → xs ++ [ v ] ++ ys ↭ v ∷ xs ++ ys
+ shift v [] ys = refl
+ shift v (x ∷ xs) ys = begin
+ x ∷ (xs ++ [ v ] ++ ys) <⟨ shift v xs ys ⟩
+ x ∷ v ∷ xs ++ ys <<⟨ refl ⟩
+ v ∷ x ∷ xs ++ ys ∎
+
+ drop-mid-≡ : ∀ {x} ws xs {ys} {zs} →
+ ws ++ [ x ] ++ ys ≡ xs ++ [ x ] ++ zs →
+ ws ++ ys ↭ xs ++ zs
+ drop-mid-≡ [] [] refl = refl
+ drop-mid-≡ [] (x ∷ xs) refl = shift _ xs _
+ drop-mid-≡ (w ∷ ws) [] refl = ↭-sym (shift _ ws _)
+ drop-mid-≡ (w ∷ ws) (x ∷ xs) eq with Lₚ.∷-injective eq
+ ... | refl , eq′ = prep w (drop-mid-≡ ws xs eq′)
+
+ drop-mid : ∀ {x} ws xs {ys zs} →
+ ws ++ [ x ] ++ ys ↭ xs ++ [ x ] ++ zs →
+ ws ++ ys ↭ xs ++ zs
+ drop-mid {x} ws xs p = drop-mid′ p ws xs refl refl
+ where
+ drop-mid′ : ∀ {l′ l″ : List A} → l′ ↭ l″ →
+ ∀ ws xs {ys zs : List A} →
+ ws ++ [ x ] ++ ys ≡ l′ →
+ xs ++ [ x ] ++ zs ≡ l″ →
+ ws ++ ys ↭ xs ++ zs
+ drop-mid′ refl ws xs refl eq = drop-mid-≡ ws xs (≡.sym eq)
+ drop-mid′ (prep x p) [] [] refl refl = p
+ drop-mid′ (prep x p) [] (x ∷ xs) refl refl = trans p (shift _ _ _)
+ drop-mid′ (prep x p) (w ∷ ws) [] refl refl = trans (↭-sym (shift _ _ _)) p
+ drop-mid′ (prep x p) (w ∷ ws) (x ∷ xs) refl refl = prep _ (drop-mid′ p ws xs refl refl)
+ drop-mid′ (swap y z p) [] [] refl refl = prep _ p
+ drop-mid′ (swap y z p) [] (x ∷ []) refl refl = prep _ p
+ drop-mid′ (swap y z p) [] (x ∷ _ ∷ xs) refl refl = prep _ (trans p (shift _ _ _))
+ drop-mid′ (swap y z p) (w ∷ []) [] refl refl = prep _ p
+ drop-mid′ (swap y z p) (w ∷ x ∷ ws) [] refl refl = prep _ (trans (↭-sym (shift _ _ _)) p)
+ drop-mid′ (swap y y p) (y ∷ []) (y ∷ []) refl refl = prep _ p
+ drop-mid′ (swap y z p) (y ∷ []) (z ∷ y ∷ xs) refl refl = begin
+ _ ∷ _ <⟨ p ⟩
+ _ ∷ (xs ++ _ ∷ _) <⟨ shift _ _ _ ⟩
+ _ ∷ _ ∷ xs ++ _ <<⟨ refl ⟩
+ _ ∷ _ ∷ xs ++ _ ∎
+ drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ []) refl refl = begin
+ _ ∷ _ ∷ ws ++ _ <<⟨ refl ⟩
+ _ ∷ (_ ∷ ws ++ _) <⟨ ↭-sym (shift _ _ _) ⟩
+ _ ∷ (ws ++ _ ∷ _) <⟨ p ⟩
+ _ ∷ _ ∎
+ drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ y ∷ xs) refl refl = swap y z (drop-mid′ p _ _ refl refl)
+ drop-mid′ (trans p₁ p₂) ws xs refl refl with ∈-∃++ _ (∈-resp-↭ p₁ (∈-insert A ws))
+ ... | (h , t , refl) = trans (drop-mid′ p₁ ws h refl refl) (drop-mid′ p₂ h xs refl refl)
+
+ -- Algebraic properties
+
+ ++-identityˡ : LeftIdentity {A = List A} _↭_ [] _++_
+ ++-identityˡ xs = refl
+
+ ++-identityʳ : RightIdentity {A = List A} _↭_ [] _++_
+ ++-identityʳ xs = ↭-reflexive (Lₚ.++-identityʳ xs)
+
+ ++-identity : Identity {A = List A} _↭_ [] _++_
+ ++-identity = ++-identityˡ , ++-identityʳ
+
+ ++-assoc : Associative {A = List A} _↭_ _++_
+ ++-assoc xs ys zs = ↭-reflexive (Lₚ.++-assoc xs ys zs)
+
+ ++-comm : Commutative _↭_ _++_
+ ++-comm [] ys = ↭-sym (++-identityʳ ys)
+ ++-comm (x ∷ xs) ys = begin
+ x ∷ xs ++ ys ↭⟨ prep x (++-comm xs ys) ⟩
+ x ∷ ys ++ xs ≡⟨ cong (λ v → x ∷ v ++ xs) (≡.sym (Lₚ.++-identityʳ _)) ⟩
+ (x ∷ ys ++ []) ++ xs ↭⟨ ++⁺ʳ xs (↭-sym (shift x ys [])) ⟩
+ (ys ++ [ x ]) ++ xs ↭⟨ ++-assoc ys [ x ] xs ⟩
+ ys ++ ([ x ] ++ xs) ≡⟨⟩
+ ys ++ (x ∷ xs) ∎
+
+ ++-isSemigroup : IsSemigroup _↭_ _++_
+ ++-isSemigroup = record
+ { isEquivalence = ↭-isEquivalence
+ ; assoc = ++-assoc
+ ; ∙-cong = ++⁺
+ }
+
+ ++-semigroup : Semigroup a _
+ ++-semigroup = record
+ { isSemigroup = ++-isSemigroup
+ }
+
+ ++-isMonoid : IsMonoid _↭_ _++_ []
+ ++-isMonoid = record
+ { isSemigroup = ++-isSemigroup
+ ; identity = ++-identity
+ }
+
+ ++-monoid : Monoid a _
+ ++-monoid = record
+ { isMonoid = ++-isMonoid
+ }
+
+ ++-isCommutativeMonoid : IsCommutativeMonoid _↭_ _++_ []
+ ++-isCommutativeMonoid = record
+ { isSemigroup = ++-isSemigroup
+ ; identityˡ = ++-identityˡ
+ ; comm = ++-comm
+ }
+
+ ++-commutativeMonoid : CommutativeMonoid _ _
+ ++-commutativeMonoid = record
+ { isCommutativeMonoid = ++-isCommutativeMonoid
+ }
+
+------------------------------------------------------------------------
+-- _∷_
+
+module _ {a} {A : Set a} where
+
+ drop-∷ : ∀ {x : A} {xs ys} → x ∷ xs ↭ x ∷ ys → xs ↭ ys
+ drop-∷ = drop-mid [] []
+
+------------------------------------------------------------------------
+-- _∷ʳ_
+
+module _ {a} {A : Set a} where
+
+ ∷↭∷ʳ : ∀ (x : A) xs → x ∷ xs ↭ xs ∷ʳ x
+ ∷↭∷ʳ x xs = ↭-sym (begin
+ xs ++ [ x ] ↭⟨ shift x xs [] ⟩
+ x ∷ xs ++ [] ≡⟨ Lₚ.++-identityʳ _ ⟩
+ x ∷ xs ∎)
+
+------------------------------------------------------------------------
+-- Relationships to other relations
+
+module _ {a} {A : Set a} where
+
+ ↭⇒~bag : _↭_ ⇒ _∼[ bag ]_
+ ↭⇒~bag xs↭ys {v} = inverse (to xs↭ys) (from xs↭ys) (from∘to xs↭ys) (to∘from xs↭ys)
+ where
+ to : ∀ {xs ys} → xs ↭ ys → v ∈ xs → v ∈ ys
+ to xs↭ys = Any-resp-↭ {A = A} xs↭ys
+
+ from : ∀ {xs ys} → xs ↭ ys → v ∈ ys → v ∈ xs
+ from xs↭ys = Any-resp-↭ (↭-sym xs↭ys)
+
+ from∘to : ∀ {xs ys} (p : xs ↭ ys) (q : v ∈ xs) → from p (to p q) ≡ q
+ from∘to refl v∈xs = refl
+ from∘to (prep _ _) (here refl) = refl
+ from∘to (prep _ p) (there v∈xs) = cong there (from∘to p v∈xs)
+ from∘to (swap x y p) (here refl) = refl
+ from∘to (swap x y p) (there (here refl)) = refl
+ from∘to (swap x y p) (there (there v∈xs)) = cong (there ∘ there) (from∘to p v∈xs)
+ from∘to (trans p₁ p₂) v∈xs
+ rewrite from∘to p₂ (Any-resp-↭ p₁ v∈xs)
+ | from∘to p₁ v∈xs = refl
+
+ to∘from : ∀ {xs ys} (p : xs ↭ ys) (q : v ∈ ys) → to p (from p q) ≡ q
+ to∘from p with from∘to (↭-sym p)
+ ... | res rewrite ↭-sym-involutive p = res
+
+ ~bag⇒↭ : _∼[ bag ]_ ⇒ _↭_
+ ~bag⇒↭ {[]} eq with empty-unique (Inv.sym eq)
+ ... | refl = refl
+ ~bag⇒↭ {x ∷ xs} eq with ∈-∃++ A (to ⟨$⟩ (here ≡.refl))
+ where open Inv.Inverse (eq {x})
+ ... | zs₁ , zs₂ , p rewrite p = begin
+ x ∷ xs <⟨ ~bag⇒↭ (drop-cons (Inv._∘_ (comm zs₁ (x ∷ zs₂)) eq)) ⟩
+ x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩
+ x ∷ (zs₁ ++ zs₂) ↭⟨ ↭-sym (shift x zs₁ zs₂) ⟩
+ zs₁ ++ x ∷ zs₂ ∎
+ where open CommutativeMonoid (commutativeMonoid bag A)
diff --git a/src/Data/List/Relation/Pointwise.agda b/src/Data/List/Relation/Pointwise.agda
new file mode 100644
index 0000000..6a454aa
--- /dev/null
+++ b/src/Data/List/Relation/Pointwise.agda
@@ -0,0 +1,273 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Pointwise lifting of relations to lists
+------------------------------------------------------------------------
+
+module Data.List.Relation.Pointwise where
+
+open import Function
+open import Function.Inverse using (Inverse)
+open import Data.Product hiding (map)
+open import Data.List.Base hiding (map ; head ; tail)
+open import Data.Fin using (Fin) renaming (zero to fzero; suc to fsuc)
+open import Data.Nat using (ℕ; zero; suc)
+open import Level
+open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec using (map′)
+open import Relation.Binary renaming (Rel to Rel₂)
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+
+infixr 5 _∷_
+
+data Pointwise {a b ℓ} {A : Set a} {B : Set b}
+ (_∼_ : REL A B ℓ) : List A → List B → Set ℓ where
+ [] : Pointwise _∼_ [] []
+ _∷_ : ∀ {x xs y ys} (x∼y : x ∼ y) (xs∼ys : Pointwise _∼_ xs ys) →
+ Pointwise _∼_ (x ∷ xs) (y ∷ ys)
+
+------------------------------------------------------------------------
+-- Operations
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ} where
+
+ head : ∀ {x y xs ys} → Pointwise _∼_ (x ∷ xs) (y ∷ ys) → x ∼ y
+ head (x∼y ∷ xs∼ys) = x∼y
+
+ tail : ∀ {x y xs ys} → Pointwise _∼_ (x ∷ xs) (y ∷ ys) →
+ Pointwise _∼_ xs ys
+ tail (x∼y ∷ xs∼ys) = xs∼ys
+
+ rec : ∀ {c} (P : ∀ {xs ys} → Pointwise _∼_ xs ys → Set c) →
+ (∀ {x y xs ys} {xs∼ys : Pointwise _∼_ xs ys} →
+ (x∼y : x ∼ y) → P xs∼ys → P (x∼y ∷ xs∼ys)) →
+ P [] →
+ ∀ {xs ys} (xs∼ys : Pointwise _∼_ xs ys) → P xs∼ys
+ rec P c n [] = n
+ rec P c n (x∼y ∷ xs∼ys) = c x∼y (rec P c n xs∼ys)
+
+ map : ∀ {ℓ₂} {_≈_ : REL A B ℓ₂} →
+ _≈_ ⇒ _∼_ → Pointwise _≈_ ⇒ Pointwise _∼_
+ map ≈⇒∼ [] = []
+ map ≈⇒∼ (x≈y ∷ xs≈ys) = ≈⇒∼ x≈y ∷ map ≈⇒∼ xs≈ys
+
+------------------------------------------------------------------------
+-- Relational properties
+
+reflexive : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b}
+ {_≈_ : REL A B ℓ₁} {_∼_ : REL A B ℓ₂} →
+ _≈_ ⇒ _∼_ → Pointwise _≈_ ⇒ Pointwise _∼_
+reflexive ≈⇒∼ [] = []
+reflexive ≈⇒∼ (x≈y ∷ xs≈ys) = ≈⇒∼ x≈y ∷ reflexive ≈⇒∼ xs≈ys
+
+refl : ∀ {a ℓ} {A : Set a} {_∼_ : Rel₂ A ℓ} →
+ Reflexive _∼_ → Reflexive (Pointwise _∼_)
+refl rfl {[]} = []
+refl rfl {x ∷ xs} = rfl ∷ refl rfl
+
+symmetric : ∀ {a b ℓ₁ ℓ₂} {A : Set a} {B : Set b}
+ {_≈_ : REL A B ℓ₁} {_∼_ : REL B A ℓ₂} →
+ Sym _≈_ _∼_ → Sym (Pointwise _≈_) (Pointwise _∼_)
+symmetric sym [] = []
+symmetric sym (x∼y ∷ xs∼ys) = sym x∼y ∷ symmetric sym xs∼ys
+
+transitive : ∀ {a b c ℓ₁ ℓ₂ ℓ₃}
+ {A : Set a} {B : Set b} {C : Set c}
+ {_≋_ : REL A B ℓ₁} {_≈_ : REL B C ℓ₂} {_∼_ : REL A C ℓ₃} →
+ Trans _≋_ _≈_ _∼_ →
+ Trans (Pointwise _≋_) (Pointwise _≈_) (Pointwise _∼_)
+transitive trans [] [] = []
+transitive trans (x∼y ∷ xs∼ys) (y∼z ∷ ys∼zs) =
+ trans x∼y y∼z ∷ transitive trans xs∼ys ys∼zs
+
+antisymmetric : ∀ {a ℓ₁ ℓ₂} {A : Set a}
+ {_≈_ : Rel₂ A ℓ₁} {_≤_ : Rel₂ A ℓ₂} →
+ Antisymmetric _≈_ _≤_ →
+ Antisymmetric (Pointwise _≈_) (Pointwise _≤_)
+antisymmetric antisym [] [] = []
+antisymmetric antisym (x∼y ∷ xs∼ys) (y∼x ∷ ys∼xs) =
+ antisym x∼y y∼x ∷ antisymmetric antisym xs∼ys ys∼xs
+
+respects₂ : ∀ {a ℓ₁ ℓ₂} {A : Set a}
+ {_≈_ : Rel₂ A ℓ₁} {_∼_ : Rel₂ A ℓ₂} →
+ _∼_ Respects₂ _≈_ →
+ (Pointwise _∼_) Respects₂ (Pointwise _≈_)
+respects₂ {_≈_ = _≈_} {_∼_} resp = resp¹ , resp²
+ where
+ resp¹ : ∀ {xs} → (Pointwise _∼_ xs) Respects (Pointwise _≈_)
+ resp¹ [] [] = []
+ resp¹ (x≈y ∷ xs≈ys) (z∼x ∷ zs∼xs) =
+ proj₁ resp x≈y z∼x ∷ resp¹ xs≈ys zs∼xs
+
+ resp² : ∀ {ys} → (flip (Pointwise _∼_) ys) Respects (Pointwise _≈_)
+ resp² [] [] = []
+ resp² (x≈y ∷ xs≈ys) (x∼z ∷ xs∼zs) =
+ proj₂ resp x≈y x∼z ∷ resp² xs≈ys xs∼zs
+
+decidable : ∀ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ} →
+ Decidable _∼_ → Decidable (Pointwise _∼_)
+decidable dec [] [] = yes []
+decidable dec [] (y ∷ ys) = no (λ ())
+decidable dec (x ∷ xs) [] = no (λ ())
+decidable dec (x ∷ xs) (y ∷ ys) with dec x y
+... | no ¬x∼y = no (¬x∼y ∘ head)
+... | yes x∼y with decidable dec xs ys
+... | no ¬xs∼ys = no (¬xs∼ys ∘ tail)
+... | yes xs∼ys = yes (x∼y ∷ xs∼ys)
+
+isEquivalence : ∀ {a ℓ} {A : Set a} {_≈_ : Rel₂ A ℓ} →
+ IsEquivalence _≈_ → IsEquivalence (Pointwise _≈_)
+isEquivalence eq = record
+ { refl = refl Eq.refl
+ ; sym = symmetric Eq.sym
+ ; trans = transitive Eq.trans
+ } where module Eq = IsEquivalence eq
+
+isPreorder : ∀ {a ℓ₁ ℓ₂} {A : Set a}
+ {_≈_ : Rel₂ A ℓ₁} {_∼_ : Rel₂ A ℓ₂} →
+ IsPreorder _≈_ _∼_ → IsPreorder (Pointwise _≈_) (Pointwise _∼_)
+isPreorder pre = record
+ { isEquivalence = isEquivalence Pre.isEquivalence
+ ; reflexive = reflexive Pre.reflexive
+ ; trans = transitive Pre.trans
+ } where module Pre = IsPreorder pre
+
+isPartialOrder : ∀ {a ℓ₁ ℓ₂} {A : Set a}
+ {_≈_ : Rel₂ A ℓ₁} {_≤_ : Rel₂ A ℓ₂} →
+ IsPartialOrder _≈_ _≤_ →
+ IsPartialOrder (Pointwise _≈_) (Pointwise _≤_)
+isPartialOrder po = record
+ { isPreorder = isPreorder PO.isPreorder
+ ; antisym = antisymmetric PO.antisym
+ } where module PO = IsPartialOrder po
+
+isDecEquivalence : ∀ {a ℓ} {A : Set a} {_≈_ : Rel₂ A ℓ} →
+ IsDecEquivalence _≈_ →
+ IsDecEquivalence (Pointwise _≈_)
+isDecEquivalence eq = record
+ { isEquivalence = isEquivalence DE.isEquivalence
+ ; _≟_ = decidable DE._≟_
+ } where module DE = IsDecEquivalence eq
+
+preorder : ∀ {p₁ p₂ p₃} → Preorder p₁ p₂ p₃ → Preorder _ _ _
+preorder p = record
+ { isPreorder = isPreorder (Preorder.isPreorder p)
+ }
+
+poset : ∀ {c ℓ₁ ℓ₂} → Poset c ℓ₁ ℓ₂ → Poset _ _ _
+poset p = record
+ { isPartialOrder = isPartialOrder (Poset.isPartialOrder p)
+ }
+
+setoid : ∀ {c ℓ} → Setoid c ℓ → Setoid _ _
+setoid s = record
+ { isEquivalence = isEquivalence (Setoid.isEquivalence s)
+ }
+
+decSetoid : ∀ {c ℓ} → DecSetoid c ℓ → DecSetoid _ _
+decSetoid d = record
+ { isDecEquivalence = isDecEquivalence (DecSetoid.isDecEquivalence d)
+ }
+
+------------------------------------------------------------------------
+-- tabulate
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ} where
+
+ tabulate⁺ : ∀ {n} {f : Fin n → A} {g : Fin n → B} →
+ (∀ i → f i ∼ g i) → Pointwise _∼_ (tabulate f) (tabulate g)
+ tabulate⁺ {zero} f∼g = []
+ tabulate⁺ {suc n} f∼g = f∼g fzero ∷ tabulate⁺ (f∼g ∘ fsuc)
+
+ tabulate⁻ : ∀ {n} {f : Fin n → A} {g : Fin n → B} →
+ Pointwise _∼_ (tabulate f) (tabulate g) → (∀ i → f i ∼ g i)
+ tabulate⁻ {zero} [] ()
+ tabulate⁻ {suc n} (x∼y ∷ xs∼ys) fzero = x∼y
+ tabulate⁻ {suc n} (x∼y ∷ xs∼ys) (fsuc i) = tabulate⁻ xs∼ys i
+
+------------------------------------------------------------------------
+-- _++_
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ} where
+
+ ++⁺ : ∀ {ws xs ys zs} → Pointwise _∼_ ws xs → Pointwise _∼_ ys zs →
+ Pointwise _∼_ (ws ++ ys) (xs ++ zs)
+ ++⁺ [] ys∼zs = ys∼zs
+ ++⁺ (w∼x ∷ ws∼xs) ys∼zs = w∼x ∷ ++⁺ ws∼xs ys∼zs
+
+------------------------------------------------------------------------
+-- concat
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ} where
+
+ concat⁺ : ∀ {xss yss} → Pointwise (Pointwise _∼_) xss yss →
+ Pointwise _∼_ (concat xss) (concat yss)
+ concat⁺ [] = []
+ concat⁺ (xs∼ys ∷ xss∼yss) = ++⁺ xs∼ys (concat⁺ xss∼yss)
+
+------------------------------------------------------------------------
+-- length
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ} where
+
+ Pointwise-length : ∀ {xs ys} → Pointwise _∼_ xs ys →
+ length xs ≡ length ys
+ Pointwise-length [] = P.refl
+ Pointwise-length (x∼y ∷ xs∼ys) = P.cong ℕ.suc (Pointwise-length xs∼ys)
+
+------------------------------------------------------------------------
+-- Properties of propositional pointwise
+
+module _ {a} {A : Set a} where
+
+ Pointwise-≡⇒≡ : Pointwise {A = A} _≡_ ⇒ _≡_
+ Pointwise-≡⇒≡ [] = P.refl
+ Pointwise-≡⇒≡ (P.refl ∷ xs∼ys) with Pointwise-≡⇒≡ xs∼ys
+ ... | P.refl = P.refl
+
+ ≡⇒Pointwise-≡ : _≡_ ⇒ Pointwise {A = A} _≡_
+ ≡⇒Pointwise-≡ P.refl = refl P.refl
+
+ Pointwise-≡↔≡ : Inverse (setoid (P.setoid A)) (P.setoid (List A))
+ Pointwise-≡↔≡ = record
+ { to = record { _⟨$⟩_ = id; cong = Pointwise-≡⇒≡ }
+ ; from = record { _⟨$⟩_ = id; cong = ≡⇒Pointwise-≡ }
+ ; inverse-of = record
+ { left-inverse-of = λ _ → refl P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+ decidable-≡ : Decidable {A = A} _≡_ → Decidable {A = List A} _≡_
+ decidable-≡ dec xs ys = Dec.map′ Pointwise-≡⇒≡ ≡⇒Pointwise-≡ (xs ≟ ys)
+ where
+ open DecSetoid (decSetoid (P.decSetoid dec))
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+Rel = Pointwise
+{-# WARNING_ON_USAGE Rel
+"Warning: Rel was deprecated in v0.15.
+Please use Pointwise instead."
+#-}
+Rel≡⇒≡ = Pointwise-≡⇒≡
+{-# WARNING_ON_USAGE Rel≡⇒≡
+"Warning: Rel≡⇒≡ was deprecated in v0.15.
+Please use Pointwise-≡⇒≡ instead."
+#-}
+≡⇒Rel≡ = ≡⇒Pointwise-≡
+{-# WARNING_ON_USAGE ≡⇒Rel≡
+"Warning: ≡⇒Rel≡ was deprecated in v0.15.
+Please use ≡⇒Pointwise-≡ instead."
+#-}
+Rel↔≡ = Pointwise-≡↔≡
+{-# WARNING_ON_USAGE Rel↔≡
+"Warning: Rel↔≡ was deprecated in v0.15.
+Please use Pointwise-≡↔≡ instead."
+#-}
diff --git a/src/Data/List/Relation/Sublist/Propositional.agda b/src/Data/List/Relation/Sublist/Propositional.agda
new file mode 100644
index 0000000..65b4b7b
--- /dev/null
+++ b/src/Data/List/Relation/Sublist/Propositional.agda
@@ -0,0 +1,48 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- An inductive definition of the sublist relation. This is commonly
+-- known as an Order Preserving Embedding (OPE).
+------------------------------------------------------------------------
+
+module Data.List.Relation.Sublist.Propositional {a} {A : Set a} where
+
+open import Data.List.Base using (List; []; _∷_; [_])
+open import Data.List.Any using (here; there)
+open import Data.List.Membership.Propositional
+open import Relation.Binary using (Rel)
+open import Relation.Binary.PropositionalEquality using (refl)
+
+------------------------------------------------------------------------
+-- Type and basic combinators
+
+infix 3 _⊆_
+
+data _⊆_ : Rel (List A) a where
+ base : [] ⊆ []
+ skip : ∀ {xs y ys} → xs ⊆ ys → xs ⊆ (y ∷ ys)
+ keep : ∀ {x xs ys} → xs ⊆ ys → (x ∷ xs) ⊆ (x ∷ ys)
+
+infix 4 []⊆_
+
+[]⊆_ : ∀ xs → [] ⊆ xs
+[]⊆ [] = base
+[]⊆ x ∷ xs = skip ([]⊆ xs)
+
+------------------------------------------------------------------------
+-- A function induced by a sublist proof
+
+lookup : ∀ {xs ys} → xs ⊆ ys → {x : A} → x ∈ xs → x ∈ ys
+lookup (skip p) v = there (lookup p v)
+lookup (keep p) (here px) = here px
+lookup (keep p) (there v) = there (lookup p v)
+lookup base ()
+
+-- Conversion between membership and proofs that a singleton is a sublist
+
+from∈ : ∀ {xs x} → x ∈ xs → [ x ] ⊆ xs
+from∈ (here refl) = keep ([]⊆ _)
+from∈ (there p) = skip (from∈ p)
+
+to∈ : ∀ {xs x} → [ x ] ⊆ xs → x ∈ xs
+to∈ p = lookup p (here refl)
diff --git a/src/Data/List/Relation/Sublist/Propositional/Properties.agda b/src/Data/List/Relation/Sublist/Propositional/Properties.agda
new file mode 100644
index 0000000..00c9f90
--- /dev/null
+++ b/src/Data/List/Relation/Sublist/Propositional/Properties.agda
@@ -0,0 +1,360 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Sublist-related properties
+------------------------------------------------------------------------
+
+module Data.List.Relation.Sublist.Propositional.Properties where
+
+open import Data.Empty
+open import Data.List.Base hiding (lookup)
+open import Data.List.Properties
+open import Data.List.Any using (here; there)
+open import Data.List.Any.Properties using (here-injective; there-injective)
+open import Data.List.Membership.Propositional
+open import Data.List.Relation.Sublist.Propositional
+open import Data.Maybe as Maybe using (nothing; just)
+open import Data.Nat.Base
+open import Data.Nat.Properties
+
+open import Function
+import Function.Bijection as Bij
+open import Function.Equivalence as Equiv using (_⇔_ ; equivalence)
+import Function.Injection as Inj
+
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as Eq hiding ([_])
+open import Relation.Nullary
+import Relation.Nullary.Decidable as D
+open import Relation.Unary as U using (Pred)
+
+------------------------------------------------------------------------
+-- The `loookup` function induced by a proof that `xs ⊆ ys` is injective
+
+module _ {a} {A : Set a} where
+
+ lookup-injective : ∀ {x : A} {xs ys} {p : xs ⊆ ys} {v w : x ∈ xs} →
+ lookup p v ≡ lookup p w → v ≡ w
+ lookup-injective {p = skip p} {v} {w} eq =
+ lookup-injective (there-injective eq)
+ lookup-injective {p = keep p} {here px} {here qx} eq =
+ cong here (≡-irrelevance _ _)
+ lookup-injective {p = keep p} {there v} {there w} eq =
+ cong there (lookup-injective (there-injective eq))
+ -- impossible cases
+ lookup-injective {p = keep p} {here px} {there w} ()
+ lookup-injective {p = keep p} {there v} {here qx} ()
+ lookup-injective {p = base} {}
+
+ []∈-irrelevant : U.Irrelevant {A = List A} ([] ⊆_)
+ []∈-irrelevant base base = refl
+ []∈-irrelevant (skip p) (skip q) = cong skip ([]∈-irrelevant p q)
+
+ [x]⊆xs↔x∈xs : ∀ {x : A} {xs} → ([ x ] ⊆ xs) Bij.⤖ (x ∈ xs)
+ [x]⊆xs↔x∈xs {x} = Bij.bijection to∈ from∈ (to∈-injective _ _) to∈∘from∈≗id
+
+ where
+
+ to∈-injective : ∀ {xs x} (p q : [ x ] ⊆ xs) → to∈ p ≡ to∈ q → p ≡ q
+ to∈-injective (skip p) (skip q) eq =
+ cong skip (to∈-injective p q (there-injective eq))
+ to∈-injective (keep p) (keep q) eq = cong keep ([]∈-irrelevant p q)
+ to∈-injective (skip p) (keep q) ()
+ to∈-injective (keep p) (skip q) ()
+
+ to∈∘from∈≗id : ∀ {xs} (p : x ∈ xs) → to∈ (from∈ p) ≡ p
+ to∈∘from∈≗id (here refl) = refl
+ to∈∘from∈≗id (there p) = cong there (to∈∘from∈≗id p)
+
+------------------------------------------------------------------------
+-- Some properties
+
+module _ {a} {A : Set a} where
+
+ ⊆-length : ∀ {xs ys : List A} → xs ⊆ ys → length xs ≤ length ys
+ ⊆-length base = ≤-refl
+ ⊆-length (skip p) = ≤-step (⊆-length p)
+ ⊆-length (keep p) = s≤s (⊆-length p)
+
+------------------------------------------------------------------------
+-- Relational properties
+
+module _ {a} {A : Set a} where
+
+ ⊆-reflexive : _≡_ ⇒ _⊆_ {A = A}
+ ⊆-reflexive {[]} refl = base
+ ⊆-reflexive {x ∷ xs} refl = keep (⊆-reflexive refl)
+
+ ⊆-refl : Reflexive {A = List A} _⊆_
+ ⊆-refl = ⊆-reflexive refl
+
+ ⊆-trans : Transitive {A = List A} _⊆_
+ ⊆-trans p base = p
+ ⊆-trans p (skip q) = skip (⊆-trans p q)
+ ⊆-trans (skip p) (keep q) = skip (⊆-trans p q)
+ ⊆-trans (keep p) (keep q) = keep (⊆-trans p q)
+
+ open ≤-Reasoning
+
+ ⊆-antisym : Antisymmetric {A = List A} _≡_ _⊆_
+ -- Positive cases
+ ⊆-antisym base base = refl
+ ⊆-antisym (keep p) (keep q) = cong (_ ∷_) (⊆-antisym p q)
+ -- Dismissing the impossible cases
+ ⊆-antisym {x ∷ xs} {y ∷ ys} (skip p) (skip q) = ⊥-elim $ 1+n≰n $ begin
+ length (y ∷ ys) ≤⟨ ⊆-length q ⟩
+ length xs ≤⟨ n≤1+n (length xs) ⟩
+ length (x ∷ xs) ≤⟨ ⊆-length p ⟩
+ length ys ∎
+ ⊆-antisym {x ∷ xs} {y ∷ ys} (skip p) (keep q) = ⊥-elim $ 1+n≰n $ begin
+ length (x ∷ xs) ≤⟨ ⊆-length p ⟩
+ length ys ≤⟨ ⊆-length q ⟩
+ length xs ∎
+ ⊆-antisym {x ∷ xs} {y ∷ ys} (keep p) (skip q) = ⊥-elim $ 1+n≰n $ begin
+ length (y ∷ ys) ≤⟨ ⊆-length q ⟩
+ length xs ≤⟨ ⊆-length p ⟩
+ length ys ∎
+
+ ⊆-minimum : Minimum (_⊆_ {A = A}) []
+ ⊆-minimum = []⊆_
+
+module _ {a} (A : Set a) where
+
+ ⊆-isPreorder : IsPreorder _≡_ (_⊆_ {A = A})
+ ⊆-isPreorder = record
+ { isEquivalence = isEquivalence
+ ; reflexive = ⊆-reflexive
+ ; trans = ⊆-trans
+ }
+
+ ⊆-preorder : Preorder _ _ _
+ ⊆-preorder = record
+ { isPreorder = ⊆-isPreorder
+ }
+
+ ⊆-isPartialOrder : IsPartialOrder _≡_ _⊆_
+ ⊆-isPartialOrder = record
+ { isPreorder = ⊆-isPreorder
+ ; antisym = ⊆-antisym
+ }
+
+ ⊆-poset : Poset _ _ _
+ ⊆-poset = record
+ { isPartialOrder = ⊆-isPartialOrder
+ }
+
+import Relation.Binary.PartialOrderReasoning as PosetReasoning
+module ⊆-Reasoning {a} {A : Set a} where
+ private module P = PosetReasoning (⊆-poset A)
+ open P public
+ renaming (_≤⟨_⟩_ to _⊆⟨_⟩_; _≈⟨⟩_ to _≡⟨⟩_)
+ hiding (_≈⟨_⟩_)
+
+
+------------------------------------------------------------------------
+-- Various functions' outputs are sublists
+
+module _ {a} {A : Set a} where
+
+ tail-⊆ : (xs : List A) → Maybe.All (_⊆ xs) (tail xs)
+ tail-⊆ [] = nothing
+ tail-⊆ (x ∷ xs) = just (skip ⊆-refl)
+
+ take-⊆ : ∀ n (xs : List A) → take n xs ⊆ xs
+ take-⊆ zero xs = []⊆ xs
+ take-⊆ (suc n) [] = []⊆ []
+ take-⊆ (suc n) (x ∷ xs) = keep (take-⊆ n xs)
+
+ drop-⊆ : ∀ n (xs : List A) → drop n xs ⊆ xs
+ drop-⊆ zero xs = ⊆-refl
+ drop-⊆ (suc n) [] = []⊆ []
+ drop-⊆ (suc n) (x ∷ xs) = skip (drop-⊆ n xs)
+
+module _ {a p} {A : Set a} {P : Pred A p} (P? : U.Decidable P) where
+
+ takeWhile-⊆ : ∀ xs → takeWhile P? xs ⊆ xs
+ takeWhile-⊆ [] = []⊆ []
+ takeWhile-⊆ (x ∷ xs) with P? x
+ ... | yes p = keep (takeWhile-⊆ xs)
+ ... | no ¬p = []⊆ x ∷ xs
+
+ dropWhile-⊆ : ∀ xs → dropWhile P? xs ⊆ xs
+ dropWhile-⊆ [] = []⊆ []
+ dropWhile-⊆ (x ∷ xs) with P? x
+ ... | yes p = skip (dropWhile-⊆ xs)
+ ... | no ¬p = ⊆-refl
+
+ filter-⊆ : ∀ xs → filter P? xs ⊆ xs
+ filter-⊆ [] = []⊆ []
+ filter-⊆ (x ∷ xs) with P? x
+ ... | yes p = keep (filter-⊆ xs)
+ ... | no ¬p = skip (filter-⊆ xs)
+
+------------------------------------------------------------------------
+-- Various functions are increasing wrt _⊆_
+
+
+------------------------------------------------------------------------
+-- _∷_
+
+module _ {a} {A : Set a} where
+
+ ∷⁻ : ∀ x {us vs : List A} → x ∷ us ⊆ x ∷ vs → us ⊆ vs
+ ∷⁻ x (skip p) = ⊆-trans (skip ⊆-refl) p
+ ∷⁻ x (keep p) = p
+
+-- map
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ map⁺ : ∀ {xs ys} (f : A → B) → xs ⊆ ys → map f xs ⊆ map f ys
+ map⁺ f base = base
+ map⁺ f (skip p) = skip (map⁺ f p)
+ map⁺ f (keep p) = keep (map⁺ f p)
+
+ open Inj.Injection
+
+ map⁻ : ∀ {xs ys} (inj : A Inj.↣ B) →
+ map (inj ⟨$⟩_) xs ⊆ map (inj ⟨$⟩_) ys → xs ⊆ ys
+ map⁻ {[]} {ys} inj p = []⊆ ys
+ map⁻ {x ∷ xs} {[]} inj ()
+ map⁻ {x ∷ xs} {y ∷ ys} inj p
+ with inj ⟨$⟩ x | inspect (inj ⟨$⟩_) x
+ | inj ⟨$⟩ y | inspect (inj ⟨$⟩_) y
+ | injective inj {x} {y}
+ map⁻ {x ∷ xs} {y ∷ ys} inj (skip p)
+ | fx | Eq.[ eq ] | fy | _ | _ = skip (map⁻ inj (coerce p))
+ where coerce = subst (λ fx → fx ∷ _ ⊆ _) (sym eq)
+ map⁻ {x ∷ xs} {y ∷ ys} inj (keep p)
+ | fx | _ | fx | _ | eq
+ rewrite eq refl = keep (map⁻ inj p)
+
+-- _++_
+
+module _ {a} {A : Set a} where
+
+ ++⁺ : ∀ {xs ys us vs : List A} → xs ⊆ ys → us ⊆ vs → xs ++ us ⊆ ys ++ vs
+ ++⁺ base q = q
+ ++⁺ (skip p) q = skip (++⁺ p q)
+ ++⁺ (keep p) q = keep (++⁺ p q)
+
+ ++⁻ : ∀ xs {us vs : List A} → xs ++ us ⊆ xs ++ vs → us ⊆ vs
+ ++⁻ [] p = p
+ ++⁻ (x ∷ xs) p = ++⁻ xs (∷⁻ x p)
+
+ skips : ∀ {xs ys} (zs : List A) → xs ⊆ ys → xs ⊆ zs ++ ys
+ skips zs = ++⁺ ([]⊆ zs)
+
+-- take / drop
+
+ ≤-take-⊆ : ∀ {m n} → m ≤ n → (xs : List A) → take m xs ⊆ take n xs
+ ≤-take-⊆ z≤n xs = []⊆ take _ xs
+ ≤-take-⊆ (s≤s p) [] = []⊆ []
+ ≤-take-⊆ (s≤s p) (x ∷ xs) = keep (≤-take-⊆ p xs)
+
+ ≥-drop-⊆ : ∀ {m n} → m ≥ n → (xs : List A) → drop m xs ⊆ drop n xs
+ ≥-drop-⊆ {m} z≤n xs = drop-⊆ m xs
+ ≥-drop-⊆ (s≤s p) [] = []⊆ []
+ ≥-drop-⊆ (s≤s p) (x ∷ xs) = ≥-drop-⊆ p xs
+
+ drop⁺ : ∀ n {xs ys : List A} → xs ⊆ ys → drop n xs ⊆ drop n ys
+ drop⁺ zero p = p
+ drop⁺ (suc n) base = []⊆ []
+ drop⁺ (suc n) (keep p) = drop⁺ n p
+ drop⁺ (suc n) {xs} {y ∷ ys} (skip p) = begin
+ drop (suc n) xs ⊆⟨ ≥-drop-⊆ (n≤1+n n) xs ⟩
+ drop n xs ⊆⟨ drop⁺ n p ⟩
+ drop n ys ∎ where open ⊆-Reasoning
+
+module _ {a p q} {A : Set a} {P : Pred A p} {Q : Pred A q}
+ (P? : U.Decidable P) (Q? : U.Decidable Q) where
+
+ ⊆-takeWhile-⊆ : (P U.⊆ Q) → ∀ xs → takeWhile P? xs ⊆ takeWhile Q? xs
+ ⊆-takeWhile-⊆ P⊆Q [] = []⊆ []
+ ⊆-takeWhile-⊆ P⊆Q (x ∷ xs) with P? x | Q? x
+ ... | yes px | yes qx = keep (⊆-takeWhile-⊆ P⊆Q xs)
+ ... | yes px | no ¬qx = ⊥-elim $ ¬qx (P⊆Q px)
+ ... | no ¬px | _ = []⊆ _
+
+ ⊇-dropWhile-⊆ : (P U.⊇ Q) → ∀ xs → dropWhile P? xs ⊆ dropWhile Q? xs
+ ⊇-dropWhile-⊆ P⊇Q [] = []⊆ []
+ ⊇-dropWhile-⊆ P⊇Q (x ∷ xs) with P? x | Q? x
+ ... | yes px | yes qx = ⊇-dropWhile-⊆ P⊇Q xs
+ ... | yes px | no ¬qx = skip (dropWhile-⊆ P? xs)
+ ... | no ¬px | yes qx = ⊥-elim $ ¬px (P⊇Q qx)
+ ... | no ¬px | no ¬qx = ⊆-refl
+
+-- filter
+
+ ⊆-filter-⊆ : (P U.⊆ Q) → ∀ xs → filter P? xs ⊆ filter Q? xs
+ ⊆-filter-⊆ P⊆Q [] = []⊆ []
+ ⊆-filter-⊆ P⊆Q (x ∷ xs) with P? x | Q? x
+ ... | yes px | yes qx = keep (⊆-filter-⊆ P⊆Q xs)
+ ... | yes px | no ¬qx = ⊥-elim $ ¬qx (P⊆Q px)
+ ... | no ¬px | yes qx = skip (⊆-filter-⊆ P⊆Q xs)
+ ... | no ¬px | no ¬qx = ⊆-filter-⊆ P⊆Q xs
+
+module _ {a p} {A : Set a} {P : Pred A p} (P? : U.Decidable P) where
+
+ filter⁺ : ∀ {xs ys : List A} → xs ⊆ ys → filter P? xs ⊆ filter P? ys
+ filter⁺ base = base
+ filter⁺ {xs} {y ∷ ys} (skip p) with P? y
+ ... | yes py = skip (filter⁺ p)
+ ... | no ¬py = filter⁺ p
+ filter⁺ {x ∷ xs} {x ∷ ys} (keep p) with P? x
+ ... | yes px = keep (filter⁺ p)
+ ... | no ¬px = filter⁺ p
+
+-- reverse
+
+module _ {a} {A : Set a} where
+
+ open ⊆-Reasoning
+
+ reverse⁺ : {xs ys : List A} → xs ⊆ ys → reverse xs ⊆ reverse ys
+ reverse⁺ base = []⊆ []
+ reverse⁺ {xs} {y ∷ ys} (skip p) = begin
+ reverse xs ≡⟨ sym $ ++-identityʳ _ ⟩
+ reverse xs ++ [] ⊆⟨ ++⁺ (reverse⁺ p) ([]⊆ _) ⟩
+ reverse ys ∷ʳ y ≡⟨ sym $ unfold-reverse y ys ⟩
+ reverse (y ∷ ys) ∎
+ reverse⁺ {x ∷ xs} {x ∷ ys} (keep p) = begin
+ reverse (x ∷ xs) ≡⟨ unfold-reverse x xs ⟩
+ reverse xs ∷ʳ x ⊆⟨ ++⁺ (reverse⁺ p) ⊆-refl ⟩
+ reverse ys ∷ʳ x ≡⟨ sym $ unfold-reverse x ys ⟩
+ reverse (x ∷ ys) ∎
+
+ reverse⁻ : {xs ys : List A} → reverse xs ⊆ reverse ys → xs ⊆ ys
+ reverse⁻ {xs} {ys} p = cast (reverse⁺ p) where
+ cast = subst₂ _⊆_ (reverse-involutive xs) (reverse-involutive ys)
+
+
+------------------------------------------------------------------------
+-- Inversion lemmas
+
+module _ {a} {A : Set a} where
+
+ keep⁻¹ : ∀ (x : A) {xs ys} → (xs ⊆ ys) ⇔ (x ∷ xs ⊆ x ∷ ys)
+ keep⁻¹ x = equivalence keep (∷⁻ x)
+
+ skip⁻¹ : ∀ {x y : A} {xs ys} → x ≢ y → (x ∷ xs ⊆ ys) ⇔ (x ∷ xs ⊆ y ∷ ys)
+ skip⁻¹ ¬eq = equivalence skip $ λ where
+ (skip p) → p
+ (keep p) → ⊥-elim (¬eq refl)
+
+ ++⁻¹ : ∀ (ps : List A) {xs ys} → (xs ⊆ ys) ⇔ (ps ++ xs ⊆ ps ++ ys)
+ ++⁻¹ ps = equivalence (++⁺ ⊆-refl) (++⁻ ps)
+
+------------------------------------------------------------------------
+-- Decidability of the order
+
+module Decidable
+ {a} {A : Set a} (eq? : Decidable {A = A} _≡_) where
+
+ infix 3 _⊆?_
+ _⊆?_ : Decidable {A = List A} _⊆_
+ [] ⊆? ys = yes ([]⊆ ys)
+ x ∷ xs ⊆? [] = no λ ()
+ x ∷ xs ⊆? y ∷ ys with eq? x y
+ ... | yes refl = D.map (keep⁻¹ x) (xs ⊆? ys)
+ ... | no ¬eq = D.map (skip⁻¹ ¬eq) (x ∷ xs ⊆? ys)
diff --git a/src/Data/List/Relation/Sublist/Propositional/Solver.agda b/src/Data/List/Relation/Sublist/Propositional/Solver.agda
new file mode 100644
index 0000000..5b1e946
--- /dev/null
+++ b/src/Data/List/Relation/Sublist/Propositional/Solver.agda
@@ -0,0 +1,144 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A solver for proving that one list is a sublist of the other.
+------------------------------------------------------------------------
+
+module Data.List.Relation.Sublist.Propositional.Solver where
+
+open import Data.Fin as Fin
+open import Data.Maybe as M
+open import Data.Nat as Nat
+open import Data.Product
+open import Data.Vec as Vec using (Vec ; lookup)
+open import Data.List hiding (lookup)
+open import Data.List.Properties
+open import Data.List.Relation.Sublist.Propositional hiding (lookup)
+open import Data.List.Relation.Sublist.Propositional.Properties
+open import Function
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality hiding ([_])
+open import Relation.Nullary
+
+open ≡-Reasoning
+
+------------------------------------------------------------------------
+-- Reified list expressions
+
+-- Basic building blocks: variables and values
+data Item {a} (n : ℕ) (A : Set a) : Set a where
+ var : Fin n → Item n A
+ val : A → Item n A
+
+-- Abstract Syntax Trees
+infixr 5 _<>_
+data TList {a} (n : ℕ) (A : Set a) : Set a where
+ It : Item n A → TList n A
+ _<>_ : TList n A → TList n A → TList n A
+ [] : TList n A
+
+-- Equivalent linearised representation
+RList : ∀ {a} n → Set a → Set a
+RList n A = List (Item n A)
+
+module _ {n a} {A : Set a} where
+
+-- Semantics
+ ⟦_⟧I : Item n A → Vec (List A) n → List A
+ ⟦ var k ⟧I ρ = lookup k ρ
+ ⟦ val a ⟧I ρ = [ a ]
+
+ ⟦_⟧T : TList n A → Vec (List A) n → List A
+ ⟦ It it ⟧T ρ = ⟦ it ⟧I ρ
+ ⟦ t <> u ⟧T ρ = ⟦ t ⟧T ρ ++ ⟦ u ⟧T ρ
+ ⟦ [] ⟧T ρ = []
+
+ ⟦_⟧R : RList n A → Vec (List A) n → List A
+ ⟦ [] ⟧R ρ = []
+ ⟦ x ∷ xs ⟧R ρ = ⟦ x ⟧I ρ ++ ⟦ xs ⟧R ρ
+
+-- Orders
+ _⊆T_ : (d e : TList n A) → Set a
+ d ⊆T e = ∀ ρ → ⟦ d ⟧T ρ ⊆ ⟦ e ⟧T ρ
+
+ _⊆R_ : (d e : RList n A) → Set a
+ d ⊆R e = ∀ ρ → ⟦ d ⟧R ρ ⊆ ⟦ e ⟧R ρ
+
+-- Flattening in a semantics-respecting manner
+
+ ⟦++⟧R : ∀ xs ys ρ → ⟦ xs ++ ys ⟧R ρ ≡ ⟦ xs ⟧R ρ ++ ⟦ ys ⟧R ρ
+ ⟦++⟧R [] ys ρ = refl
+ ⟦++⟧R (x ∷ xs) ys ρ = begin
+ ⟦ x ⟧I ρ ++ ⟦ xs ++ ys ⟧R ρ
+ ≡⟨ cong (⟦ x ⟧I ρ ++_) (⟦++⟧R xs ys ρ) ⟩
+ ⟦ x ⟧I ρ ++ ⟦ xs ⟧R ρ ++ ⟦ ys ⟧R ρ
+ ≡⟨ sym $ ++-assoc (⟦ x ⟧I ρ) (⟦ xs ⟧R ρ) (⟦ ys ⟧R ρ) ⟩
+ (⟦ x ⟧I ρ ++ ⟦ xs ⟧R ρ) ++ ⟦ ys ⟧R ρ
+ ∎
+
+ flatten : ∀ (t : TList n A) → Σ[ r ∈ RList n A ] ⟦ t ⟧T ≗ ⟦ r ⟧R
+ flatten [] = [] , λ _ → refl
+ flatten (It it) = it ∷ [] , λ ρ → sym $ ++-identityʳ (⟦ It it ⟧T ρ)
+ flatten (t <> u) =
+ let (rt , eqt) = flatten t
+ (ru , equ) = flatten u
+ in rt ++ ru , λ ρ → begin
+ ⟦ t <> u ⟧T ρ ≡⟨⟩
+ ⟦ t ⟧T ρ ++ ⟦ u ⟧T ρ ≡⟨ cong₂ _++_ (eqt ρ) (equ ρ) ⟩
+ ⟦ rt ⟧R ρ ++ ⟦ ru ⟧R ρ ≡⟨ sym $ ⟦++⟧R rt ru ρ ⟩
+ ⟦ rt ++ ru ⟧R ρ ∎
+
+------------------------------------------------------------------------
+-- Solver for the sublist problem
+
+module _ {n : ℕ} {a} {A : Set a} (eq? : Decidable {A = A} _≡_) where
+
+-- auxiliary lemmas
+
+ private
+
+ keep-it : ∀ it (xs ys : RList n A) → xs ⊆R ys → (it ∷ xs) ⊆R (it ∷ ys)
+ keep-it it xs ys hyp ρ = ++⁺ ⊆-refl (hyp ρ)
+
+ skip-it : ∀ it (d e : RList n A) → d ⊆R e → d ⊆R (it ∷ e)
+ skip-it it d ys hyp ρ = skips (⟦ it ⟧I ρ) (hyp ρ)
+
+-- Solver for linearised expressions
+
+ solveR : (d e : RList n A) → Maybe (d ⊆R e)
+ -- trivial
+ solveR [] e = just (λ ρ → []⊆ ⟦ e ⟧R ρ)
+ solveR (it ∷ d) [] = nothing
+ -- actual work
+ solveR (var k ∷ d) (var l ∷ e) with k Fin.≟ l
+ ... | yes refl = M.map (keep-it (var k) d e) (solveR d e)
+ ... | no ¬eq = M.map (skip-it (var l) (var k ∷ d) e) (solveR (var k ∷ d) e)
+ solveR (val a ∷ d) (val b ∷ e) with eq? a b
+ ... | yes refl = M.map (keep-it (val a) d e) (solveR d e)
+ ... | no ¬eq = M.map (skip-it (val b) (val a ∷ d) e) (solveR (val a ∷ d) e)
+ solveR d (x ∷ e) = M.map (skip-it x d e) (solveR d e)
+
+-- Coming back to ASTs thanks to flatten
+
+ solveT : (t u : TList n A) → Maybe (t ⊆T u)
+ solveT t u =
+ let (rt , eqt) = flatten t
+ (ru , equ) = flatten u
+ in case solveR rt ru of λ where
+ (just hyp) → just (λ ρ → subst₂ _⊆_ (sym (eqt ρ)) (sym (equ ρ)) (hyp ρ))
+ nothing → nothing
+
+-- Prover for ASTs
+
+ prove : ∀ d e → From-just (solveT d e)
+ prove d e = from-just (solveT d e)
+
+------------------------------------------------------------------------
+-- Test
+
+_ : ∀ xs ys → xs ++ xs ⊆ (xs ++ 2 ∷ ys) ++ xs
+_ = λ xs ys →
+ let `xs = It (var zero)
+ `ys = It (var (suc zero))
+ in prove Nat._≟_ (`xs <> `xs) ((`xs <> It (val 2) <> `ys) <> `xs)
+ (Vec.fromList (xs ∷ ys ∷ []))
diff --git a/src/Data/List/Relation/Subset/Propositional.agda b/src/Data/List/Relation/Subset/Propositional.agda
new file mode 100644
index 0000000..9fc1a99
--- /dev/null
+++ b/src/Data/List/Relation/Subset/Propositional.agda
@@ -0,0 +1,16 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- The sublist relation over propositional equality.
+------------------------------------------------------------------------
+
+module Data.List.Relation.Subset.Propositional
+ {a} {A : Set a} where
+
+import Data.List.Relation.Subset.Setoid as SetoidSubset
+open import Relation.Binary.PropositionalEquality using (setoid)
+
+------------------------------------------------------------------------
+-- Re-export parameterised definitions from setoid sublists
+
+open SetoidSubset (setoid A) public
diff --git a/src/Data/List/Relation/Subset/Propositional/Properties.agda b/src/Data/List/Relation/Subset/Propositional/Properties.agda
new file mode 100644
index 0000000..3eb927e
--- /dev/null
+++ b/src/Data/List/Relation/Subset/Propositional/Properties.agda
@@ -0,0 +1,208 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of the sublist relation over setoid equality.
+------------------------------------------------------------------------
+
+open import Relation.Binary hiding (Decidable)
+
+module Data.List.Relation.Subset.Propositional.Properties
+ where
+
+open import Category.Monad
+open import Data.Bool.Base using (Bool; true; false; T)
+open import Data.List
+open import Data.List.Any using (Any; here; there)
+open import Data.List.Any.Properties
+open import Data.List.Categorical
+open import Data.List.Membership.Propositional
+open import Data.List.Membership.Propositional.Properties
+import Data.List.Relation.Subset.Setoid.Properties as Setoidₚ
+open import Data.List.Relation.Subset.Propositional
+import Data.Product as Prod
+import Data.Sum as Sum
+open import Function using (_∘_; _∘′_; id; _$_)
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Inverse as Inv using (_↔_; module Inverse)
+open import Function.Equivalence using (module Equivalence)
+open import Relation.Nullary using (¬_; yes; no)
+open import Relation.Unary using (Decidable)
+open import Relation.Binary using (_⇒_)
+open import Relation.Binary.PropositionalEquality
+ using (_≡_; _≗_; isEquivalence; refl; setoid; module ≡-Reasoning)
+import Relation.Binary.PreorderReasoning as PreorderReasoning
+
+private
+ open module ListMonad {ℓ} = RawMonad (monad {ℓ = ℓ})
+
+------------------------------------------------------------------------
+-- Relational properties
+
+module _ {a} {A : Set a} where
+
+ ⊆-reflexive : _≡_ ⇒ (_⊆_ {A = A})
+ ⊆-reflexive refl = id
+
+ ⊆-refl : Reflexive {A = List A} _⊆_
+ ⊆-refl x∈xs = x∈xs
+
+ ⊆-trans : Transitive {A = List A} (_⊆_ {A = A})
+ ⊆-trans xs⊆ys ys⊆zs x∈xs = ys⊆zs (xs⊆ys x∈xs)
+
+module _ {a} (A : Set a) where
+
+ ⊆-isPreorder : IsPreorder {A = List A} _≡_ _⊆_
+ ⊆-isPreorder = record
+ { isEquivalence = isEquivalence
+ ; reflexive = ⊆-reflexive
+ ; trans = ⊆-trans
+ }
+
+ ⊆-preorder : Preorder _ _ _
+ ⊆-preorder = record
+ { isPreorder = ⊆-isPreorder
+ }
+
+------------------------------------------------------------------------
+-- Reasoning over subsets
+
+module ⊆-Reasoning {a} (A : Set a) where
+ open Setoidₚ.⊆-Reasoning public
+ hiding (_≋⟨_⟩_) renaming (_≋⟨⟩_ to _≡⟨⟩_)
+
+------------------------------------------------------------------------
+-- Properties relating _⊆_ to various list functions
+------------------------------------------------------------------------
+-- Any
+
+module _ {a p} {A : Set a} {P : A → Set p} {xs ys : List A} where
+
+ mono : 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
+
+module _ {a b} {A : Set a} {B : Set b} (f : A → B) {xs ys} where
+
+ map-mono : xs ⊆ ys → map f xs ⊆ map f ys
+ map-mono xs⊆ys =
+ _⟨$⟩_ (Inverse.to map-∈↔) ∘
+ Prod.map id (Prod.map xs⊆ys id) ∘
+ _⟨$⟩_ (Inverse.from map-∈↔)
+
+------------------------------------------------------------------------
+-- _++_
+
+module _ {a} {A : Set a} {xs₁ xs₂ ys₁ ys₂ : List A} where
+
+ _++-mono_ : xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → xs₁ ++ xs₂ ⊆ ys₁ ++ ys₂
+ _++-mono_ xs₁⊆ys₁ xs₂⊆ys₂ =
+ _⟨$⟩_ (Inverse.to ++↔) ∘
+ Sum.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
+ _⟨$⟩_ (Inverse.from ++↔)
+
+------------------------------------------------------------------------
+-- concat
+
+module _ {a} {A : Set a} {xss yss : List (List A)} where
+
+ concat-mono : xss ⊆ yss → concat xss ⊆ concat yss
+ concat-mono xss⊆yss =
+ _⟨$⟩_ (Inverse.to $ concat-∈↔ {a = a}) ∘
+ Prod.map id (Prod.map id xss⊆yss) ∘
+ _⟨$⟩_ (Inverse.from $ concat-∈↔ {a = a})
+
+------------------------------------------------------------------------
+-- _>>=_
+
+module _ {ℓ} {A B : Set ℓ} (f g : A → List B) {xs ys} where
+
+ >>=-mono : xs ⊆ ys → (∀ {x} → f x ⊆ g x) → (xs >>= f) ⊆ (ys >>= g)
+ >>=-mono xs⊆ys f⊆g =
+ _⟨$⟩_ (Inverse.to $ >>=-∈↔ {ℓ = ℓ}) ∘
+ Prod.map id (Prod.map xs⊆ys f⊆g) ∘
+ _⟨$⟩_ (Inverse.from $ >>=-∈↔ {ℓ = ℓ})
+
+------------------------------------------------------------------------
+-- _⊛_
+
+module _ {ℓ} {A B : Set ℓ} {fs gs : List (A → B)} {xs ys : List A} where
+
+ _⊛-mono_ : fs ⊆ gs → xs ⊆ ys → (fs ⊛ xs) ⊆ (gs ⊛ ys)
+ _⊛-mono_ 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)
+
+------------------------------------------------------------------------
+-- _⊗_
+
+module _ {ℓ} {A B : Set ℓ} {xs₁ ys₁ : List A} {xs₂ ys₂ : List B} where
+
+ _⊗-mono_ : xs₁ ⊆ ys₁ → xs₂ ⊆ ys₂ → (xs₁ ⊗ xs₂) ⊆ (ys₁ ⊗ ys₂)
+ xs₁⊆ys₁ ⊗-mono xs₂⊆ys₂ =
+ _⟨$⟩_ (Inverse.to $ ⊗-∈↔ {ℓ = ℓ}) ∘
+ Prod.map xs₁⊆ys₁ xs₂⊆ys₂ ∘
+ _⟨$⟩_ (Inverse.from $ ⊗-∈↔ {ℓ = ℓ})
+
+------------------------------------------------------------------------
+-- any
+
+module _ {a} {A : Set a} (p : A → Bool) {xs ys} where
+
+ any-mono : xs ⊆ ys → T (any p xs) → T (any p ys)
+ any-mono xs⊆ys =
+ _⟨$⟩_ (Equivalence.to any⇔) ∘
+ mono xs⊆ys ∘
+ _⟨$⟩_ (Equivalence.from any⇔)
+
+------------------------------------------------------------------------
+-- map-with-∈
+
+module _ {a b} {A : Set a} {B : Set b}
+ {xs : List A} {f : ∀ {x} → x ∈ xs → B}
+ {ys : List A} {g : ∀ {x} → x ∈ ys → B} where
+
+ map-with-∈-mono : (xs⊆ys : xs ⊆ ys) → (∀ {x} → f {x} ≗ g ∘ xs⊆ys) →
+ map-with-∈ xs f ⊆ map-with-∈ ys g
+ map-with-∈-mono 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 ≡-Reasoning
+
+------------------------------------------------------------------------
+-- filter
+
+module _ {a p} {A : Set a} {P : A → Set p} (P? : Decidable P) where
+
+ filter⁺ : ∀ xs → filter P? xs ⊆ xs
+ filter⁺ = Setoidₚ.filter⁺ (setoid A) P?
+
+------------------------------------------------------------------------
+-- DEPRECATED
+------------------------------------------------------------------------
+
+-- Version 0.16
+
+boolFilter-⊆ : ∀ {a} {A : Set a} (p : A → Bool) →
+ (xs : List A) → boolFilter p xs ⊆ xs
+boolFilter-⊆ _ [] = λ ()
+boolFilter-⊆ p (x ∷ xs) with p x | boolFilter-⊆ p xs
+... | false | hyp = there ∘ hyp
+... | true | hyp =
+ λ { (here eq) → here eq
+ ; (there ∈boolFilter) → there (hyp ∈boolFilter)
+ }
+{-# WARNING_ON_USAGE boolFilter-⊆
+"Warning: boolFilter was deprecated in v0.15.
+Please use filter instead."
+#-}
diff --git a/src/Data/List/Relation/Subset/Setoid.agda b/src/Data/List/Relation/Subset/Setoid.agda
new file mode 100644
index 0000000..86ef526
--- /dev/null
+++ b/src/Data/List/Relation/Subset/Setoid.agda
@@ -0,0 +1,28 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- The extensional sublist relation over setoid equality.
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Data.List.Relation.Subset.Setoid
+ {c ℓ} (S : Setoid c ℓ) where
+
+open import Data.List using (List)
+open import Data.List.Membership.Setoid S using (_∈_)
+open import Level using (_⊔_)
+open import Relation.Nullary using (¬_)
+
+open Setoid S renaming (Carrier to A)
+
+------------------------------------------------------------------------
+-- Definitions
+
+infix 4 _⊆_ _⊈_
+
+_⊆_ : Rel (List A) (c ⊔ ℓ)
+xs ⊆ ys = ∀ {x} → x ∈ xs → x ∈ ys
+
+_⊈_ : Rel (List A) (c ⊔ ℓ)
+xs ⊈ ys = ¬ xs ⊆ ys
diff --git a/src/Data/List/Relation/Subset/Setoid/Properties.agda b/src/Data/List/Relation/Subset/Setoid/Properties.agda
new file mode 100644
index 0000000..abd866d
--- /dev/null
+++ b/src/Data/List/Relation/Subset/Setoid/Properties.agda
@@ -0,0 +1,78 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of the extensional sublist relation over setoid equality.
+------------------------------------------------------------------------
+
+open import Relation.Binary hiding (Decidable)
+
+module Data.List.Relation.Subset.Setoid.Properties where
+
+open import Data.Bool using (Bool; true; false)
+open import Data.List
+open import Data.List.Any using (here; there)
+import Data.List.Membership.Setoid as Membership
+open import Data.List.Membership.Setoid.Properties
+import Data.List.Relation.Subset.Setoid as Sublist
+import Data.List.Relation.Equality.Setoid as Equality
+open import Relation.Nullary using (¬_; yes; no)
+open import Relation.Unary using (Pred; Decidable)
+import Relation.Binary.PreorderReasoning as PreorderReasoning
+
+open Setoid using (Carrier)
+
+------------------------------------------------------------------------
+-- Relational properties
+
+module _ {a ℓ} (S : Setoid a ℓ) where
+
+ open Equality S
+ open Sublist S
+ open Membership S
+
+ ⊆-reflexive : _≋_ ⇒ _⊆_
+ ⊆-reflexive xs≋ys = ∈-resp-≋ S xs≋ys
+
+ ⊆-refl : Reflexive _⊆_
+ ⊆-refl x∈xs = x∈xs
+
+ ⊆-trans : Transitive _⊆_
+ ⊆-trans xs⊆ys ys⊆zs x∈xs = ys⊆zs (xs⊆ys x∈xs)
+
+ ⊆-isPreorder : IsPreorder _≋_ _⊆_
+ ⊆-isPreorder = record
+ { isEquivalence = ≋-isEquivalence
+ ; reflexive = ⊆-reflexive
+ ; trans = ⊆-trans
+ }
+
+ ⊆-preorder : Preorder _ _ _
+ ⊆-preorder = record
+ { isPreorder = ⊆-isPreorder
+ }
+
+ -- Reasoning over subsets
+ module ⊆-Reasoning where
+ open PreorderReasoning ⊆-preorder public
+ renaming (_∼⟨_⟩_ to _⊆⟨_⟩_ ; _≈⟨_⟩_ to _≋⟨_⟩_ ; _≈⟨⟩_ to _≋⟨⟩_)
+
+ infix 1 _∈⟨_⟩_
+ _∈⟨_⟩_ : ∀ x {xs ys} → x ∈ xs → xs IsRelatedTo ys → x ∈ ys
+ x ∈⟨ x∈xs ⟩ xs⊆ys = (begin xs⊆ys) x∈xs
+
+------------------------------------------------------------------------
+-- filter
+
+module _ {a p ℓ} (S : Setoid a ℓ)
+ {P : Pred (Carrier S) p} (P? : Decidable P) where
+
+ open Setoid S renaming (Carrier to A)
+ open Sublist S
+
+ filter⁺ : ∀ xs → filter P? xs ⊆ xs
+ filter⁺ [] ()
+ filter⁺ (x ∷ xs) y∈f[x∷xs] with P? x
+ ... | no _ = there (filter⁺ xs y∈f[x∷xs])
+ ... | yes _ with y∈f[x∷xs]
+ ... | here y≈x = here y≈x
+ ... | there y∈f[xs] = there (filter⁺ xs y∈f[xs])
diff --git a/src/Data/List/Reverse.agda b/src/Data/List/Reverse.agda
index 40c5296..451f516 100644
--- a/src/Data/List/Reverse.agda
+++ b/src/Data/List/Reverse.agda
@@ -6,10 +6,9 @@
module Data.List.Reverse where
-open import Data.List.Base
-open import Data.Nat
-import Data.Nat.Properties as Nat
-open import Induction.Nat using (<′-Rec; <′-rec)
+open import Data.List.Base as L hiding (reverse)
+open import Data.List.Properties
+open import Function
open import Relation.Binary.PropositionalEquality
-- If you want to traverse a list from the end, then you can use the
@@ -17,25 +16,17 @@ open import Relation.Binary.PropositionalEquality
infixl 5 _∶_∶ʳ_
-data Reverse {A : Set} : List A → Set where
+data Reverse {a} {A : Set a} : List A → Set a where
[] : Reverse []
_∶_∶ʳ_ : ∀ 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
- where
- Pred : ℕ → Set
- Pred n = (xs : List A) → length xs ≡ n → Reverse xs
-
- lem : ∀ xs {x : A} → length xs <′ length (xs ∷ʳ x)
- lem [] = ≤′-refl
- lem (x ∷ xs) = Nat.s≤′s (lem xs)
-
- 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
- with rec (length xs) (lem xs) xs refl
- rev ._ rec .([] ∷ʳ x) refl | ._ ∷ʳ' x | [] = _ ∶ [] ∶ʳ x
- rev ._ rec .(ys ∷ʳ y ∷ʳ x) refl | ._ ∷ʳ' x | ys ∶ rs ∶ʳ y =
- _ ∶ (_ ∶ rs ∶ʳ y) ∶ʳ x
+module _ {a} {A : Set a} where
+
+ reverse : (xs : List A) → Reverse (L.reverse xs)
+ reverse [] = []
+ reverse (x ∷ xs) = cast $ _ ∶ reverse xs ∶ʳ x where
+ cast = subst Reverse (sym $ unfold-reverse x xs)
+
+ reverseView : (xs : List A) → Reverse xs
+ reverseView xs = cast $ reverse (L.reverse xs) where
+ cast = subst Reverse (reverse-involutive xs)
diff --git a/src/Data/List/Solver.agda b/src/Data/List/Solver.agda
new file mode 100644
index 0000000..d7414b2
--- /dev/null
+++ b/src/Data/List/Solver.agda
@@ -0,0 +1,19 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Automatic solvers for equations over lists
+------------------------------------------------------------------------
+
+-- See README.Nat for examples of how to use similar solvers
+
+module Data.List.Solver where
+
+import Algebra.Solver.Monoid as Solver
+open import Data.List.Properties using (++-monoid)
+
+------------------------------------------------------------------------
+-- A module for automatically solving propositional equivalences
+-- containing _++_
+
+module ++-Solver {a} {A : Set a} =
+ Solver (++-monoid A) renaming (id to nil)
diff --git a/src/Data/List/Zipper.agda b/src/Data/List/Zipper.agda
new file mode 100644
index 0000000..63c2389
--- /dev/null
+++ b/src/Data/List/Zipper.agda
@@ -0,0 +1,116 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- List Zippers, basic types and operations
+------------------------------------------------------------------------
+
+module Data.List.Zipper where
+
+open import Data.Nat.Base
+open import Data.Maybe.Base as Maybe using (Maybe ; just ; nothing)
+open import Data.List.Base as List using (List ; [] ; _∷_)
+open import Function
+
+
+-- Definition
+------------------------------------------------------------------------
+
+-- A List Zipper represents a List together with a particular sub-List
+-- in focus. The user can attempt to move the focus left or right, with
+-- a risk of failure if one has already reached the corresponding end.
+
+-- To make these operations efficient, the `context` the sub List in
+-- focus lives in is stored *backwards*. This is made formal by `toList`
+-- which returns the List a Zipper represents.
+
+record Zipper {a} (A : Set a) : Set a where
+ constructor mkZipper
+ field context : List A
+ value : List A
+
+ toList : List A
+ toList = List.reverse context List.++ value
+open Zipper public
+
+-- Embedding Lists as Zippers without any context
+fromList : ∀ {a} {A : Set a} → List A → Zipper A
+fromList = mkZipper []
+
+-- Fundamental operations of a Zipper: Moving around
+------------------------------------------------------------------------
+
+module _ {a} {A : Set a} where
+
+ left : Zipper A → Maybe (Zipper A)
+ left (mkZipper [] val) = nothing
+ left (mkZipper (x ∷ ctx) val) = just (mkZipper ctx (x ∷ val))
+
+ right : Zipper A → Maybe (Zipper A)
+ right (mkZipper ctx []) = nothing
+ right (mkZipper ctx (x ∷ val)) = just (mkZipper (x ∷ ctx) val)
+
+
+-- Focus-respecting operations
+------------------------------------------------------------------------
+
+module _ {a} {A : Set a} where
+
+ reverse : Zipper A → Zipper A
+ reverse (mkZipper ctx val) = mkZipper val ctx
+
+ -- If we think of a List [x₁⋯xₘ] split into a List [xₙ₊₁⋯xₘ] in focus
+ -- of another list [x₁⋯xₙ] then there are 4 places (marked {k} here) in
+ -- which we can insert new values: [{1}x₁⋯xₙ{2}][{3}xₙ₊₁⋯xₘ{4}]
+
+ -- The following 4 functions implement these 4 insertions.
+
+ -- `xs ˢ++ zp` inserts `xs` on the `s` side of the context of the Zipper `zp`
+ -- `zp ++ˢ xs` insert `xs` on the `s` side of the value in focus of the Zipper `zp`
+
+ infixr 5 _ˡ++_ _ʳ++_
+ infixl 5 _++ˡ_ _++ʳ_
+ -- {1}
+ _ˡ++_ : List A → Zipper A → Zipper A
+ xs ˡ++ mkZipper ctx val = mkZipper (ctx List.++ List.reverse xs) val
+
+ -- {2}
+ _ʳ++_ : List A → Zipper A → Zipper A
+ xs ʳ++ mkZipper ctx val = mkZipper (List.reverse xs List.++ ctx) val
+
+ -- {3}
+ _++ˡ_ : Zipper A → List A → Zipper A
+ mkZipper ctx val ++ˡ xs = mkZipper ctx (xs List.++ val)
+
+ -- {4}
+ _++ʳ_ : Zipper A → List A → Zipper A
+ mkZipper ctx val ++ʳ xs = mkZipper ctx (val List.++ xs)
+
+
+-- List-like operations
+------------------------------------------------------------------------
+
+module _ {a} {A : Set a} where
+
+ length : Zipper A → ℕ
+ length (mkZipper ctx val) = List.length ctx + List.length val
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ map : (A → B) → Zipper A → Zipper B
+ map f (mkZipper ctx val) = (mkZipper on List.map f) ctx val
+
+ foldr : (A → B → B) → B → Zipper A → B
+ foldr c n (mkZipper ctx val) = List.foldl (flip c) (List.foldr c n val) ctx
+
+
+-- Generating all the possible foci of a list
+------------------------------------------------------------------------
+
+module _ {a} {A : Set a} where
+
+ allFociIn : List A → List A → List (Zipper A)
+ allFociIn ctx [] = List.[ mkZipper ctx [] ]
+ allFociIn ctx xxs@(x ∷ xs) = mkZipper ctx xxs ∷ allFociIn (x ∷ ctx) xs
+
+ allFoci : List A → List (Zipper A)
+ allFoci = allFociIn []
diff --git a/src/Data/List/Zipper/Properties.agda b/src/Data/List/Zipper/Properties.agda
new file mode 100644
index 0000000..0f58ff2
--- /dev/null
+++ b/src/Data/List/Zipper/Properties.agda
@@ -0,0 +1,131 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- List Zipper-related properties
+------------------------------------------------------------------------
+
+module Data.List.Zipper.Properties where
+
+open import Data.Maybe.Base as Maybe using (Maybe ; just ; nothing)
+open import Data.List.Base as List using (List ; [] ; _∷_)
+open import Data.List.Properties
+open import Data.List.Zipper
+open import Relation.Binary.PropositionalEquality
+open ≡-Reasoning
+open import Function
+
+-- Invariant: Zipper represents a given list
+------------------------------------------------------------------------
+
+module _ {a} {A : Set a} where
+
+ -- Stability under moving left or right
+
+ toList-left-identity : (zp : Zipper A) → Maybe.All ((_≡_ on toList) zp) (left zp)
+ toList-left-identity (mkZipper [] val) = nothing
+ toList-left-identity (mkZipper (x ∷ ctx) val) = just $′ begin
+ List.reverse (x ∷ ctx) List.++ val
+ ≡⟨ cong (List._++ val) (unfold-reverse x ctx) ⟩
+ (List.reverse ctx List.++ List.[ x ]) List.++ val
+ ≡⟨ ++-assoc (List.reverse ctx) List.[ x ] val ⟩
+ toList (mkZipper ctx (x ∷ val))
+ ∎
+
+ toList-right-identity : (zp : Zipper A) → Maybe.All ((_≡_ on toList) zp) (right zp)
+ toList-right-identity (mkZipper ctx []) = nothing
+ toList-right-identity (mkZipper ctx (x ∷ val)) = just $′ begin
+ List.reverse ctx List.++ x ∷ val
+ ≡⟨ sym (++-assoc (List.reverse ctx) List.[ x ] val) ⟩
+ (List.reverse ctx List.++ List.[ x ]) List.++ val
+ ≡⟨ cong (List._++ val) (sym (unfold-reverse x ctx)) ⟩
+ List.reverse (x ∷ ctx) List.++ val
+ ∎
+
+ -- Applying reverse does correspond to reversing the represented list
+
+ toList-reverse-commute : (zp : Zipper A) → toList (reverse zp) ≡ List.reverse (toList zp)
+ toList-reverse-commute (mkZipper ctx val) = begin
+ List.reverse val List.++ ctx
+ ≡⟨ cong (List.reverse val List.++_) (sym (reverse-involutive ctx)) ⟩
+ List.reverse val List.++ List.reverse (List.reverse ctx)
+ ≡⟨ sym (reverse-++-commute (List.reverse ctx) val) ⟩
+ List.reverse (List.reverse ctx List.++ val)
+ ∎
+
+
+-- Properties of the insertion functions
+------------------------------------------------------------------------
+
+ -- _ˡ++_ properties
+
+ toList-ˡ++-commute : ∀ xs (zp : Zipper A) → toList (xs ˡ++ zp) ≡ xs List.++ toList zp
+ toList-ˡ++-commute xs (mkZipper ctx val) = begin
+ List.reverse (ctx List.++ List.reverse xs) List.++ val
+ ≡⟨ cong (List._++ _) (reverse-++-commute ctx (List.reverse xs)) ⟩
+ (List.reverse (List.reverse xs) List.++ List.reverse ctx) List.++ val
+ ≡⟨ ++-assoc (List.reverse (List.reverse xs)) (List.reverse ctx) val ⟩
+ List.reverse (List.reverse xs) List.++ List.reverse ctx List.++ val
+ ≡⟨ cong (List._++ _) (reverse-involutive xs) ⟩
+ xs List.++ List.reverse ctx List.++ val
+ ∎
+
+ ˡ++-assoc : ∀ xs ys (zp : Zipper A) → xs ˡ++ (ys ˡ++ zp) ≡ (xs List.++ ys) ˡ++ zp
+ ˡ++-assoc xs ys (mkZipper ctx val) = cong (λ ctx → mkZipper ctx val) $ begin
+ (ctx List.++ List.reverse ys) List.++ List.reverse xs
+ ≡⟨ ++-assoc ctx _ _ ⟩
+ ctx List.++ List.reverse ys List.++ List.reverse xs
+ ≡⟨ cong (ctx List.++_) (sym (reverse-++-commute xs ys)) ⟩
+ ctx List.++ List.reverse (xs List.++ ys)
+ ∎
+
+ -- _ʳ++_ properties
+
+ ʳ++-assoc : ∀ xs ys (zp : Zipper A) → xs ʳ++ (ys ʳ++ zp) ≡ (ys List.++ xs) ʳ++ zp
+ ʳ++-assoc xs ys (mkZipper ctx val) = cong (λ ctx → mkZipper ctx val) $ begin
+ List.reverse xs List.++ List.reverse ys List.++ ctx
+ ≡⟨ sym (++-assoc (List.reverse xs) (List.reverse ys) ctx) ⟩
+ (List.reverse xs List.++ List.reverse ys) List.++ ctx
+ ≡⟨ cong (List._++ ctx) (sym (reverse-++-commute ys xs)) ⟩
+ List.reverse (ys List.++ xs) List.++ ctx
+ ∎
+
+ -- _++ˡ_ properties
+
+ ++ˡ-assoc : ∀ xs ys (zp : Zipper A) → zp ++ˡ xs ++ˡ ys ≡ zp ++ˡ (ys List.++ xs)
+ ++ˡ-assoc xs ys (mkZipper ctx val) = cong (mkZipper ctx) $ sym $ ++-assoc ys xs val
+
+ -- _++ʳ_ properties
+
+ toList-++ʳ-commute : ∀ (zp : Zipper A) xs → toList (zp ++ʳ xs) ≡ toList zp List.++ xs
+ toList-++ʳ-commute (mkZipper ctx val) xs = begin
+ List.reverse ctx List.++ val List.++ xs
+ ≡⟨ sym (++-assoc (List.reverse ctx) val xs) ⟩
+ (List.reverse ctx List.++ val) List.++ xs
+ ∎
+
+ ++ʳ-assoc : ∀ xs ys (zp : Zipper A) → zp ++ʳ xs ++ʳ ys ≡ zp ++ʳ (xs List.++ ys)
+ ++ʳ-assoc xs ys (mkZipper ctx val) = cong (mkZipper ctx) $ ++-assoc val xs ys
+
+
+-- List-like operations indeed correspond to their counterparts
+------------------------------------------------------------------------
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ toList-map-commute : ∀ (f : A → B) zp → toList (map f zp) ≡ List.map f (toList zp)
+ toList-map-commute f zp@(mkZipper ctx val) = begin
+ List.reverse (List.map f ctx) List.++ List.map f val
+ ≡⟨ cong (List._++ _) (sym (reverse-map-commute f ctx)) ⟩
+ List.map f (List.reverse ctx) List.++ List.map f val
+ ≡⟨ sym (map-++-commute f (List.reverse ctx) val) ⟩
+ List.map f (List.reverse ctx List.++ val)
+ ∎
+
+ toList-foldr-commute : ∀ (c : A → B → B) n zp → foldr c n zp ≡ List.foldr c n (toList zp)
+ toList-foldr-commute c n (mkZipper ctx val) = begin
+ List.foldl (flip c) (List.foldr c n val) ctx
+ ≡⟨ sym (reverse-foldr c (List.foldr c n val) ctx) ⟩
+ List.foldr c (List.foldr c n val) (List.reverse ctx)
+ ≡⟨ sym (foldr-++ c n (List.reverse ctx) val) ⟩
+ List.foldr c n (List.reverse ctx List.++ val)
+ ∎
diff --git a/src/Data/M.agda b/src/Data/M.agda
deleted file mode 100644
index ed77f79..0000000
--- a/src/Data/M.agda
+++ /dev/null
@@ -1,25 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- M-types (the dual of W-types)
-------------------------------------------------------------------------
-
-module Data.M where
-
-open import Level
-open import Coinduction
-
--- The family of M-types.
-
-data M {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
- inf : (x : A) (f : B x → ∞ (M A B)) → M A B
-
--- Projections.
-
-head : ∀ {a b} {A : Set a} {B : A → Set b} →
- M A B → A
-head (inf x f) = x
-
-tail : ∀ {a b} {A : Set a} {B : A → Set b} →
- (x : M A B) → B (head x) → M A B
-tail (inf x f) b = ♭ (f b)
diff --git a/src/Data/M/Indexed.agda b/src/Data/M/Indexed.agda
deleted file mode 100644
index efd87f9..0000000
--- a/src/Data/M/Indexed.agda
+++ /dev/null
@@ -1,42 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- Indexed M-types (the dual of indexed W-types aka Petersson-Synek
--- trees).
-------------------------------------------------------------------------
-
-module Data.M.Indexed where
-
-open import Level
-open import Coinduction
-open import Data.Product
-open import Data.Container.Indexed.Core
-open import Function
-open import Relation.Unary
-
--- The family of indexed M-types.
-
-module _ {ℓ c r} {O : Set ℓ} (C : Container O O c r) where
-
- data M (o : O) : Set (ℓ ⊔ c ⊔ r) where
- inf : ⟦ C ⟧ (∞ ∘ M) o → M o
-
- open Container C
-
- -- Projections.
-
- head : M ⊆ Command
- head (inf (c , _)) = c
-
- tail : ∀ {o} (m : M o) (r : Response (head m)) → M (next (head m) r)
- tail (inf (_ , k)) r = ♭ (k r)
-
- force : M ⊆ ⟦ C ⟧ M
- force (inf (c , k)) = c , λ r → ♭ (k r)
-
- -- Coiteration.
-
- coit : ∀ {ℓ} {X : Pred O ℓ} → X ⊆ ⟦ C ⟧ X → X ⊆ M
- coit ψ x = inf (proj₁ cs , λ r → ♯ coit ψ (proj₂ cs r))
- where
- cs = ψ x
diff --git a/src/Data/Maybe.agda b/src/Data/Maybe.agda
index df29181..ad3c1d5 100644
--- a/src/Data/Maybe.agda
+++ b/src/Data/Maybe.agda
@@ -6,9 +6,6 @@
module Data.Maybe where
-open import Category.Functor
-open import Category.Monad
-open import Category.Monad.Identity
open import Function
open import Level
open import Relation.Nullary
@@ -22,41 +19,6 @@ open import Relation.Unary as U
open import Data.Maybe.Base public
------------------------------------------------------------------------
--- Maybe monad
-
-functor : ∀ {f} → RawFunctor (Maybe {a = f})
-functor = record
- { _<$>_ = map
- }
-
-monadT : ∀ {f} {M : Set f → Set f} →
- RawMonad M → RawMonad (λ A → M (Maybe A))
-monadT M = record
- { return = M.return ∘ just
- ; _>>=_ = λ m f → M._>>=_ m (maybe f (M.return nothing))
- }
- where module M = RawMonad M
-
-monad : ∀ {f} → RawMonad (Maybe {a = f})
-monad = monadT IdentityMonad
-
-monadZero : ∀ {f} → RawMonadZero (Maybe {a = f})
-monadZero = record
- { monad = monad
- ; ∅ = nothing
- }
-
-monadPlus : ∀ {f} → RawMonadPlus (Maybe {a = f})
-monadPlus {f} = record
- { monadZero = monadZero
- ; _∣_ = _∣_
- }
- where
- _∣_ : {A : Set f} → Maybe A → Maybe A → Maybe A
- nothing ∣ y = y
- just x ∣ y = just x
-
-------------------------------------------------------------------------
-- Equality
data Eq {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) : Rel (Maybe A) (a ⊔ ℓ) where
diff --git a/src/Data/Maybe/Base.agda b/src/Data/Maybe/Base.agda
index 33317e5..b399444 100644
--- a/src/Data/Maybe/Base.agda
+++ b/src/Data/Maybe/Base.agda
@@ -17,12 +17,17 @@ data Maybe {a} (A : Set a) : Set a where
{-# FOREIGN GHC type AgdaMaybe a b = Maybe b #-}
{-# COMPILE GHC Maybe = data MAlonzo.Code.Data.Maybe.Base.AgdaMaybe (Just | Nothing) #-}
+open import Function
+open import Agda.Builtin.Equality using (_≡_ ; refl)
+
+just-injective : ∀ {a} {A : Set a} {a b} → (Maybe A ∋ just a) ≡ just b → a ≡ b
+just-injective refl = refl
+
------------------------------------------------------------------------
-- Some operations
open import Data.Bool.Base using (Bool; true; false; not)
open import Data.Unit.Base using (⊤)
-open import Function
open import Relation.Nullary
boolToMaybe : Bool → Maybe ⊤
@@ -52,16 +57,23 @@ maybe j n nothing = n
maybe′ : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → B → Maybe A → B
maybe′ = maybe
+-- A defaulting mechanism
+
+fromMaybe : ∀ {a} {A : Set a} → A → Maybe A → A
+fromMaybe = maybe′ id
+
-- A safe variant of "fromJust". If the value is nothing, then the
-- return type is the unit type.
-From-just : ∀ {a} (A : Set a) → Maybe A → Set a
-From-just A (just _) = A
-From-just A nothing = Lift ⊤
+module _ {a} {A : Set a} where
-from-just : ∀ {a} {A : Set a} (x : Maybe A) → From-just A x
-from-just (just x) = x
-from-just nothing = _
+ From-just : Maybe A → Set a
+ From-just (just _) = A
+ From-just nothing = Lift a ⊤
+
+ from-just : (x : Maybe A) → From-just x
+ from-just (just x) = x
+ from-just nothing = _
-- Functoriality: map.
@@ -93,3 +105,29 @@ to-witness (just {x = p} _) = p
to-witness-T : ∀ {p} {P : Set p} (m : Maybe P) → T (is-just m) → P
to-witness-T (just p) _ = p
to-witness-T nothing ()
+
+------------------------------------------------------------------------
+-- Aligning and Zipping
+
+open import Data.These using (These; this; that; these)
+open import Data.Product hiding (zip)
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
+
+ alignWith : (These A B → C) → Maybe A → Maybe B → Maybe C
+ alignWith f (just a) (just b) = just (f (these a b))
+ alignWith f (just a) nothing = just (f (this a))
+ alignWith f nothing (just b) = just (f (that b))
+ alignWith f nothing nothing = nothing
+
+ zipWith : (A → B → C) → Maybe A → Maybe B → Maybe C
+ zipWith f (just a) (just b) = just (f a b)
+ zipWith _ _ _ = nothing
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ align : Maybe A → Maybe B → Maybe (These A B)
+ align = alignWith id
+
+ zip : Maybe A → Maybe B → Maybe (A × B)
+ zip = zipWith _,_
diff --git a/src/Data/Maybe/Categorical.agda b/src/Data/Maybe/Categorical.agda
new file mode 100644
index 0000000..e3e6f99
--- /dev/null
+++ b/src/Data/Maybe/Categorical.agda
@@ -0,0 +1,81 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A categorical view of Maybe
+------------------------------------------------------------------------
+
+module Data.Maybe.Categorical where
+
+open import Data.Maybe.Base
+open import Category.Functor
+open import Category.Applicative
+open import Category.Monad
+import Function.Identity.Categorical as Id
+open import Function
+
+------------------------------------------------------------------------
+-- Maybe applicative functor
+
+functor : ∀ {f} → RawFunctor {f} Maybe
+functor = record
+ { _<$>_ = map
+ }
+
+applicative : ∀ {f} → RawApplicative {f} Maybe
+applicative = record
+ { pure = just
+ ; _⊛_ = maybe map (const nothing)
+ }
+
+------------------------------------------------------------------------
+-- Maybe monad transformer
+
+monadT : ∀ {f} → RawMonadT {f} (_∘′ Maybe)
+monadT M = record
+ { return = M.return ∘ just
+ ; _>>=_ = λ m f → m M.>>= maybe f (M.return nothing)
+ }
+ where module M = RawMonad M
+
+------------------------------------------------------------------------
+-- Maybe monad
+
+monad : ∀ {f} → RawMonad {f} Maybe
+monad = monadT Id.monad
+
+monadZero : ∀ {f} → RawMonadZero {f} Maybe
+monadZero = record
+ { monad = monad
+ ; ∅ = nothing
+ }
+
+monadPlus : ∀ {f} → RawMonadPlus {f} Maybe
+monadPlus {f} = record
+ { monadZero = monadZero
+ ; _∣_ = maybe′ (const ∘ just) id
+ }
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {f F} (App : RawApplicative {f} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → Maybe (F A) → F (Maybe A)
+ sequenceA nothing = pure nothing
+ sequenceA (just x) = just <$> x
+
+ mapA : ∀ {a} {A : Set a} {B} → (A → F B) → Maybe A → F (Maybe B)
+ mapA f = sequenceA ∘ map f
+
+ forA : ∀ {a} {A : Set a} {B} → Maybe A → (A → F B) → F (Maybe B)
+ forA = flip mapA
+
+module _ {m M} (Mon : RawMonad {m} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM = sequenceA App
+ mapM = mapA App
+ forM = forA App
diff --git a/src/Data/Nat.agda b/src/Data/Nat.agda
index 522478e..4d860dc 100644
--- a/src/Data/Nat.agda
+++ b/src/Data/Nat.agda
@@ -6,23 +6,26 @@
module Data.Nat where
-open import Function
-open import Function.Equality as F using (_⟨$⟩_)
-open import Function.Injection using (_↣_)
-open import Data.Sum
-open import Data.Empty
-import Level
-open import Relation.Nullary
-import Relation.Nullary.Decidable as Dec
-open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_; refl)
-
--- The core type and operations are re-exported from Data.Nat.Base
+------------------------------------------------------------------------
+-- Publicly re-export the contents of the base module
+
open import Data.Nat.Base public
--- If there is an injection from a type to ℕ, then the type has
--- decidable equality.
+------------------------------------------------------------------------
+-- Publicly re-export queries
+
+open import Data.Nat.Properties public
+ using
+ ( _≟_
+ ; _≤?_ ; _≥?_ ; _<?_ ; _>?_
+ ; _≤′?_; _≥′?_; _<′?_; _>′?_
+ ; _≤″?_; _<″?_; _≥″?_; _>″?_
+ )
+
+------------------------------------------------------------------------
+-- Deprecated
+
+-- Version 0.17
-eq? : ∀ {a} {A : Set a} → A ↣ ℕ → Decidable {A = A} _≡_
-eq? inj = Dec.via-injection inj _≟_
+open import Data.Nat.Properties public
+ using (≤-pred)
diff --git a/src/Data/Nat/Base.agda b/src/Data/Nat/Base.agda
index 52367e1..2136fcf 100644
--- a/src/Data/Nat/Base.agda
+++ b/src/Data/Nat/Base.agda
@@ -6,92 +6,56 @@
module Data.Nat.Base where
-import Level using (zero)
-open import Function using (_∘_)
+open import Level using (0ℓ)
+open import Function using (_∘_; flip)
open import Relation.Binary
open import Relation.Binary.Core
open import Relation.Binary.PropositionalEquality.Core
-import Relation.Binary.PropositionalEquality.TrustMe as TrustMe
-open import Relation.Nullary using (¬_; Dec; yes; no)
-
-infix 4 _≤_ _<_ _≥_ _>_ _≰_ _≮_ _≱_ _≯_
+open import Relation.Binary.PropositionalEquality
+open import Relation.Nullary using (¬_)
------------------------------------------------------------------------
-- The types
open import Agda.Builtin.Nat public
- using ( zero; suc; _+_; _*_ )
- renaming ( Nat to ℕ
- ; _-_ to _∸_ )
+ using (zero; suc) renaming (Nat to ℕ)
-data _≤_ : Rel ℕ Level.zero where
+------------------------------------------------------------------------
+-- Standard ordering relations
+
+infix 4 _≤_ _<_ _≥_ _>_ _≰_ _≮_ _≱_ _≯_
+
+data _≤_ : Rel ℕ 0ℓ where
z≤n : ∀ {n} → zero ≤ n
s≤s : ∀ {m n} (m≤n : m ≤ n) → suc m ≤ suc n
-_<_ : Rel ℕ Level.zero
+_<_ : Rel ℕ 0ℓ
m < n = suc m ≤ n
-_≥_ : Rel ℕ Level.zero
+_≥_ : Rel ℕ 0ℓ
m ≥ n = n ≤ m
-_>_ : Rel ℕ Level.zero
+_>_ : Rel ℕ 0ℓ
m > n = n < m
-_≰_ : Rel ℕ Level.zero
+_≰_ : Rel ℕ 0ℓ
a ≰ b = ¬ a ≤ b
-_≮_ : Rel ℕ Level.zero
+_≮_ : Rel ℕ 0ℓ
a ≮ b = ¬ a < b
-_≱_ : Rel ℕ Level.zero
+_≱_ : Rel ℕ 0ℓ
a ≱ b = ¬ a ≥ b
-_≯_ : Rel ℕ Level.zero
+_≯_ : Rel ℕ 0ℓ
a ≯ b = ¬ a > b
--- The following, alternative definition of _≤_ is more suitable for
--- well-founded induction (see Induction.Nat).
-
-infix 4 _≤′_ _<′_ _≥′_ _>′_
-
-data _≤′_ (m : ℕ) : ℕ → Set where
- ≤′-refl : m ≤′ m
- ≤′-step : ∀ {n} (m≤′n : m ≤′ n) → m ≤′ suc n
-
-_<′_ : Rel ℕ Level.zero
-m <′ n = suc m ≤′ n
-
-_≥′_ : Rel ℕ Level.zero
-m ≥′ n = n ≤′ m
-
-_>′_ : Rel ℕ Level.zero
-m >′ n = n <′ m
-
--- Another alternative definition of _≤_.
-
-record _≤″_ (m n : ℕ) : Set where
- constructor less-than-or-equal
- field
- {k} : ℕ
- proof : m + k ≡ n
-
-infix 4 _≤″_ _<″_ _≥″_ _>″_
-
-_<″_ : Rel ℕ Level.zero
-m <″ n = suc m ≤″ n
-
-_≥″_ : Rel ℕ Level.zero
-m ≥″ n = n ≤″ m
-
-_>″_ : Rel ℕ Level.zero
-m >″ n = n <″ m
-
-erase : ∀ {m n} → m ≤″ n → m ≤″ n
-erase (less-than-or-equal eq) = less-than-or-equal (TrustMe.erase eq)
-
------------------------------------------------------------------------
-- Arithmetic
+open import Agda.Builtin.Nat public
+ using (_+_; _*_ ) renaming (_-_ to _∸_)
+
pred : ℕ → ℕ
pred zero = zero
pred (suc n) = n
@@ -137,33 +101,57 @@ _^_ : ℕ → ℕ → ℕ
x ^ zero = 1
x ^ suc n = x * x ^ n
+-- Distance
+
+∣_-_∣ : ℕ → ℕ → ℕ
+∣ zero - y ∣ = y
+∣ x - zero ∣ = x
+∣ suc x - suc y ∣ = ∣ x - y ∣
+
------------------------------------------------------------------------
--- Queries
+-- The following, alternative definition of _≤_ is more suitable for
+-- well-founded induction (see Induction.Nat).
-infix 4 _≟_ _≤?_
+infix 4 _≤′_ _<′_ _≥′_ _>′_
-_≟_ : Decidable {A = ℕ} _≡_
-zero ≟ zero = yes refl
-suc m ≟ suc n with m ≟ n
-suc m ≟ suc .m | yes refl = yes refl
-suc m ≟ suc n | no prf = no (prf ∘ (λ p → subst (λ x → m ≡ pred x) p refl))
-zero ≟ suc n = no λ()
-suc m ≟ zero = no λ()
+data _≤′_ (m : ℕ) : ℕ → Set where
+ ≤′-refl : m ≤′ m
+ ≤′-step : ∀ {n} (m≤′n : m ≤′ n) → m ≤′ suc n
-≤-pred : ∀ {m n} → suc m ≤ suc n → m ≤ n
-≤-pred (s≤s m≤n) = m≤n
+_<′_ : Rel ℕ 0ℓ
+m <′ n = suc m ≤′ n
-_≤?_ : Decidable _≤_
-zero ≤? _ = yes z≤n
-suc m ≤? zero = no λ()
-suc m ≤? suc n with m ≤? n
-... | yes m≤n = yes (s≤s m≤n)
-... | no m≰n = no (m≰n ∘ ≤-pred)
+_≥′_ : Rel ℕ 0ℓ
+m ≥′ n = n ≤′ m
+
+_>′_ : Rel ℕ 0ℓ
+m >′ n = n <′ m
+------------------------------------------------------------------------
+-- Another alternative definition of _≤_.
+
+record _≤″_ (m n : ℕ) : Set where
+ constructor less-than-or-equal
+ field
+ {k} : ℕ
+ proof : m + k ≡ n
+
+infix 4 _≤″_ _<″_ _≥″_ _>″_
+
+_<″_ : Rel ℕ 0ℓ
+m <″ n = suc m ≤″ n
+
+_≥″_ : Rel ℕ 0ℓ
+m ≥″ n = n ≤″ m
+
+_>″_ : Rel ℕ 0ℓ
+m >″ n = n <″ m
+
+------------------------------------------------------------------------
-- A comparison view. Taken from "View from the left"
-- (McBride/McKinna); details may differ.
-data Ordering : Rel ℕ Level.zero where
+data Ordering : Rel ℕ 0ℓ where
less : ∀ m k → Ordering m (suc (m + k))
equal : ∀ m → Ordering m m
greater : ∀ m k → Ordering (suc (m + k)) m
@@ -173,6 +161,6 @@ compare zero zero = equal zero
compare (suc m) zero = greater zero m
compare zero (suc n) = less zero n
compare (suc m) (suc n) with compare m n
-compare (suc .m) (suc .(suc m + k)) | less m k = less (suc m) k
-compare (suc .m) (suc .m) | equal m = equal (suc m)
-compare (suc .(suc m + k)) (suc .m) | greater m k = greater (suc m) k
+... | less m k = less (suc m) k
+... | equal m = equal (suc m)
+... | greater n k = greater (suc n) k
diff --git a/src/Data/Nat/Coprimality.agda b/src/Data/Nat/Coprimality.agda
index 85ed06c..19fdbf5 100644
--- a/src/Data/Nat/Coprimality.agda
+++ b/src/Data/Nat/Coprimality.agda
@@ -8,27 +8,25 @@ module Data.Nat.Coprimality where
open import Data.Empty
open import Data.Fin using (toℕ; fromℕ≤)
-open import Data.Fin.Properties as FinProp
+open import Data.Fin.Properties using (toℕ-fromℕ≤)
open import Data.Nat
-open import Data.Nat.Primality
-import Data.Nat.Properties as NatProp
-open import Data.Nat.Divisibility as Div
+open import Data.Nat.Divisibility
open import Data.Nat.GCD
open import Data.Nat.GCD.Lemmas
+open import Data.Nat.Primality
+open import Data.Nat.Properties
open import Data.Product as Prod
open import Function
-open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_; _≢_; refl)
+open import Level using (0ℓ)
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≢_; refl; cong; subst; module ≡-Reasoning)
open import Relation.Nullary
open import Relation.Binary
-open import Algebra
-private
- module P = Poset Div.poset
-- Coprime m n is inhabited iff m and n are coprime (relatively
-- prime), i.e. if their only common divisor is 1.
-Coprime : (m n : ℕ) → Set
+Coprime : Rel ℕ 0ℓ
Coprime m n = ∀ {i} → i ∣ m × i ∣ n → i ≡ 1
-- Coprime numbers have 1 as their gcd.
@@ -37,15 +35,14 @@ coprime-gcd : ∀ {m n} → Coprime m n → GCD m n 1
coprime-gcd {m} {n} c = GCD.is (1∣ m , 1∣ n) greatest
where
greatest : ∀ {d} → d ∣ m × d ∣ n → d ∣ 1
- greatest cd with c cd
- greatest .{1} cd | refl = 1∣ 1
+ greatest cd with c cd
+ ... | refl = 1∣ 1
-- If two numbers have 1 as their gcd, then they are coprime.
gcd-coprime : ∀ {m n} → GCD m n 1 → Coprime m n
gcd-coprime g cd with GCD.greatest g cd
-gcd-coprime g cd | divides q eq =
- NatProp.i*j≡1⇒j≡1 q _ (PropEq.sym eq)
+... | divides q eq = i*j≡1⇒j≡1 q _ (P.sym eq)
-- Coprime is decidable.
@@ -59,12 +56,12 @@ private
coprime? : Decidable Coprime
coprime? i j with gcd i j
... | (0 , g) = no (0≢1 ∘ GCD.unique g ∘ coprime-gcd)
-... | (1 , g) = yes (λ {i} → gcd-coprime g {i})
+... | (1 , g) = yes (gcd-coprime g)
... | (suc (suc d) , g) = no (2+≢1 ∘ GCD.unique g ∘ coprime-gcd)
-- The coprimality relation is symmetric.
-sym : ∀ {m n} → Coprime m n → Coprime n m
+sym : Symmetric Coprime
sym c = c ∘ swap
-- Everything is coprime to 1.
@@ -75,12 +72,12 @@ sym c = c ∘ swap
-- Nothing except for 1 is coprime to 0.
0-coprimeTo-1 : ∀ {m} → Coprime 0 m → m ≡ 1
-0-coprimeTo-1 {m} c = c (m ∣0 , P.refl)
+0-coprimeTo-1 {m} c = c (m ∣0 , ∣-refl)
-- If m and n are coprime, then n + m and n are also coprime.
coprime-+ : ∀ {m n} → Coprime m n → Coprime (n + m) n
-coprime-+ c (d₁ , d₂) = c (∣-∸ d₁ d₂ , d₂)
+coprime-+ c (d₁ , d₂) = c (∣m+n∣m⇒∣n d₁ d₂ , d₂)
-- If the "gcd" in Bézout's identity is non-zero, then the "other"
-- divisors are coprime.
@@ -125,14 +122,13 @@ 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 | NatProp.*-comm 0 q₁
- | q₂ * 0 | NatProp.*-comm 0 q₂
-... | .0 | refl | .0 | refl = gcd-* 1 1 (1-coprimeTo 1)
+ with q₁ * 0 | *-comm 0 q₁ | q₂ * 0 | *-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))
gcd′-gcd : ∀ {m n d} → GCD′ m n d → GCD m n d
-gcd′-gcd (gcd-* q₁ q₂ c) = GCD.is (∣-* q₁ , ∣-* q₂) (coprime-factors c)
+gcd′-gcd (gcd-* q₁ q₂ c) = GCD.is (n∣m*n q₁ , n∣m*n q₂) (coprime-factors c)
-- Calculates (the alternative representation of) the gcd of the
-- arguments.
@@ -149,11 +145,11 @@ prime⇒coprime 1 () _ _ _ _
prime⇒coprime (suc (suc m)) _ 0 () _ _
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
+ ⊥-elim $ i+1+j≢i 0 (begin
2 + m ≡⟨ 2+m≡q*0 ⟩
- q * 0 ≡⟨ proj₂ NatProp.*-zero q ⟩
+ q * 0 ≡⟨ *-zeroʳ q ⟩
0 ∎)
- where open PropEq.≡-Reasoning
+ where open ≡-Reasoning
prime⇒coprime (suc (suc m)) p (suc n) _ 1+n<2+m {suc (suc i)}
(2+i∣2+m , 2+i∣1+n) =
⊥-elim (p _ 2+i′∣2+m)
@@ -163,10 +159,9 @@ 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 NatProp.≤-Reasoning
+ where open ≤-Reasoning
2+i′∣2+m : 2 + toℕ (fromℕ≤ i<m) ∣ 2 + m
- 2+i′∣2+m = PropEq.subst
- (λ j → j ∣ 2 + m)
- (PropEq.sym (PropEq.cong (_+_ 2) (FinProp.toℕ-fromℕ≤ i<m)))
+ 2+i′∣2+m = subst (_∣ 2 + m)
+ (P.sym (cong (2 +_) (toℕ-fromℕ≤ i<m)))
2+i∣2+m
diff --git a/src/Data/Nat/DivMod.agda b/src/Data/Nat/DivMod.agda
index a07726f..746e97e 100644
--- a/src/Data/Nat/DivMod.agda
+++ b/src/Data/Nat/DivMod.agda
@@ -6,143 +6,102 @@
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
-open import Relation.Nullary.Decidable
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
-import Relation.Binary.PropositionalEquality.TrustMe as TrustMe
- using (erase)
-
-open import Agda.Builtin.Nat using ( div-helper; mod-helper )
+open import Agda.Builtin.Nat using (div-helper; mod-helper)
-open SemiringSolver
-open P.≡-Reasoning
-open ≤-Reasoning
- renaming (begin_ to start_; _∎ to _□; _≡⟨_⟩_ to _≡⟨_⟩′_)
+open import Data.Fin using (Fin; toℕ; fromℕ≤)
+open import Data.Fin.Properties using (toℕ-fromℕ≤)
+open import Data.Nat as Nat
+open import Data.Nat.DivMod.Core
+open import Data.Nat.Properties using (≤⇒≤″; +-assoc; +-comm; +-identityʳ)
+open import Function using (_$_)
+open import Relation.Nullary.Decidable using (False)
+open import Relation.Binary.PropositionalEquality
-infixl 7 _div_ _mod_ _divMod_
+open ≡-Reasoning
--- A specification of integer division.
+------------------------------------------------------------------------
+-- Basic operations
-record DivMod (dividend divisor : ℕ) : Set where
- constructor result
- field
- quotient : ℕ
- remainder : Fin divisor
- property : dividend ≡ toℕ remainder + quotient * divisor
+infixl 7 _div_ _%_
--- Integer division.
+-- Integer division
_div_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} → ℕ
-(d div 0) {}
-(d div suc s) = div-helper 0 s d s
-
--- The remainder after integer division.
+(a div 0) {}
+(a div suc n) = div-helper 0 n a n
-private
- -- The remainder is not too large.
+-- Integer remainder (mod)
- mod-lemma : (acc d n : ℕ) →
- let s = acc + n in
- mod-helper acc s d n ≤ s
+_%_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} → ℕ
+(a % 0) {}
+(a % suc n) = mod-helper 0 n a n
- mod-lemma acc zero n = start
- acc ≤⟨ m≤m+n acc n ⟩
- acc + n □
+------------------------------------------------------------------------
+-- Properties
+
+a≡a%n+[a/n]*n : ∀ a n → a ≡ a % suc n + (a div (suc n)) * suc n
+a≡a%n+[a/n]*n a n = division-lemma 0 0 a n
+
+a%1≡0 : ∀ a → a % 1 ≡ 0
+a%1≡0 = a[modₕ]1≡0
+
+a%n<n : ∀ a n → a % suc n < suc n
+a%n<n a n = s≤s (a[modₕ]n<n 0 a n)
+
+n%n≡0 : ∀ n → suc n % suc n ≡ 0
+n%n≡0 n = n[modₕ]n≡0 0 n
+
+a%n%n≡a%n : ∀ a n → a % suc n % suc n ≡ a % suc n
+a%n%n≡a%n a n = modₕ-idem 0 a n
+
+[a+n]%n≡a%n : ∀ a n → (a + suc n) % suc n ≡ a % suc n
+[a+n]%n≡a%n a n = a+n[modₕ]n≡a[modₕ]n 0 a n
+
+[a+kn]%n≡a%n : ∀ a k n → (a + k * (suc n)) % suc n ≡ a % suc n
+[a+kn]%n≡a%n a zero n = cong (_% suc n) (+-identityʳ a)
+[a+kn]%n≡a%n a (suc k) n = begin
+ (a + (m + k * m)) % m ≡⟨ cong (_% m) (sym (+-assoc a m (k * m))) ⟩
+ (a + m + k * m) % m ≡⟨ [a+kn]%n≡a%n (a + m) k n ⟩
+ (a + m) % m ≡⟨ [a+n]%n≡a%n a n ⟩
+ a % m ∎
+ where m = suc n
+
+kn%n≡0 : ∀ k n → k * (suc n) % suc n ≡ 0
+kn%n≡0 = [a+kn]%n≡a%n 0
+
+%-distribˡ-+ : ∀ a b n → (a + b) % suc n ≡ (a % suc n + b % suc n) % suc n
+%-distribˡ-+ a b n = begin
+ (a + b) % m ≡⟨ cong (λ v → (v + b) % m) (a≡a%n+[a/n]*n a n) ⟩
+ (a % m + a div m * m + b) % m ≡⟨ cong (_% m) (+-assoc (a % m) _ b) ⟩
+ (a % m + (a div m * m + b)) % m ≡⟨ cong (λ v → (a % m + v) % m) (+-comm _ b) ⟩
+ (a % m + (b + a div m * m)) % m ≡⟨ cong (_% m) (sym (+-assoc (a % m) b _)) ⟩
+ (a % m + b + a div m * m) % m ≡⟨ [a+kn]%n≡a%n (a % m + b) (a div m) n ⟩
+ (a % m + b) % m ≡⟨ cong (λ v → (a % m + v) % m) (a≡a%n+[a/n]*n b n) ⟩
+ (a % m + (b % m + (b div m) * m)) % m ≡⟨ sym (cong (_% m) (+-assoc (a % m) (b % m) _)) ⟩
+ (a % m + b % m + (b div m) * m) % m ≡⟨ [a+kn]%n≡a%n (a % m + b % m) (b div m) n ⟩
+ (a % m + b % m) % m ∎
+ where m = suc n
- mod-lemma acc (suc d) zero = start
- mod-helper zero (acc + 0) d (acc + 0) ≤⟨ mod-lemma zero d (acc + 0) ⟩
- acc + 0 □
+------------------------------------------------------------------------
+-- A specification of integer division.
- mod-lemma acc (suc d) (suc n) =
- P.subst (λ x → mod-helper (suc acc) x d n ≤ x)
- (P.sym (+-suc acc n))
- (mod-lemma (suc acc) d n)
+record DivMod (dividend divisor : ℕ) : Set where
+ constructor result
+ field
+ quotient : ℕ
+ remainder : Fin divisor
+ property : dividend ≡ toℕ remainder + quotient * divisor
_mod_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} → Fin divisor
-(d mod 0) {}
-(d mod suc s) =
- Fin.fromℕ≤″ (mod-helper 0 s d s)
- (Nat.erase (≤⇒≤″ (s≤s (mod-lemma 0 d s))))
-
--- Integer division with remainder.
-
-private
-
- -- The quotient and remainder are related to the dividend and
- -- divisor in the right way.
-
- division-lemma :
- (mod-acc div-acc d n : ℕ) →
- let s = mod-acc + n in
- mod-acc + div-acc * suc s + d
- ≡
- mod-helper mod-acc s d n + div-helper div-acc s d n * suc s
-
- division-lemma mod-acc div-acc zero n = begin
-
- mod-acc + div-acc * suc s + zero ≡⟨ +-identityʳ _ ⟩
-
- mod-acc + div-acc * suc s ∎
-
- where s = mod-acc + n
-
- division-lemma mod-acc div-acc (suc d) zero = begin
-
- mod-acc + div-acc * suc s + suc d ≡⟨ solve 3
- (λ mod-acc div-acc d →
- let s = mod-acc :+ con 0 in
- mod-acc :+ div-acc :* (con 1 :+ s) :+ (con 1 :+ d)
- :=
- (con 1 :+ div-acc) :* (con 1 :+ s) :+ d)
- P.refl mod-acc div-acc d ⟩
- suc div-acc * suc s + d ≡⟨ division-lemma zero (suc div-acc) d s ⟩
-
- mod-helper zero s d s +
- div-helper (suc div-acc) s d s * suc s ≡⟨⟩
-
- mod-helper mod-acc s (suc d) zero +
- div-helper div-acc s (suc d) zero * suc s ∎
-
- where s = mod-acc + 0
-
- division-lemma mod-acc div-acc (suc d) (suc n) = begin
-
- mod-acc + div-acc * suc s + suc d ≡⟨ solve 4
- (λ mod-acc div-acc n d →
- mod-acc :+ div-acc :* (con 1 :+ (mod-acc :+ (con 1 :+ n))) :+ (con 1 :+ d)
- :=
- con 1 :+ mod-acc :+ div-acc :* (con 2 :+ mod-acc :+ n) :+ d)
- P.refl mod-acc div-acc n d ⟩
- suc mod-acc + div-acc * suc s′ + d ≡⟨ division-lemma (suc mod-acc) div-acc d n ⟩
-
- mod-helper (suc mod-acc) s′ d n +
- div-helper div-acc s′ d n * suc s′ ≡⟨ P.cong (λ s → mod-helper (suc mod-acc) s d n +
- div-helper div-acc s d n * suc s)
- (P.sym (+-suc mod-acc n)) ⟩
-
- mod-helper (suc mod-acc) s d n +
- div-helper div-acc s d n * suc s ≡⟨⟩
-
- mod-helper mod-acc s (suc d) (suc n) +
- div-helper div-acc s (suc d) (suc n) * suc s ∎
-
- where
- s = mod-acc + suc n
- s′ = suc mod-acc + n
+(a mod 0) {}
+(a mod suc n) = fromℕ≤ (a%n<n a n)
_divMod_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} →
DivMod dividend divisor
-(d divMod 0) {}
-(d divMod suc s) =
- result (d div suc s) (d mod suc s) (TrustMe.erase (begin
- d ≡⟨ division-lemma 0 0 d s ⟩
- mod-helper 0 s d s + div-helper 0 s d s * suc s ≡⟨ P.cong₂ _+_ (P.sym (FinP.toℕ-fromℕ≤ lemma)) P.refl ⟩
- toℕ (Fin.fromℕ≤ lemma) + div-helper 0 s d s * suc s ≡⟨ P.cong (λ n → toℕ n + div-helper 0 s d s * suc s)
- (FinP.fromℕ≤≡fromℕ≤″ lemma _) ⟩
- toℕ (Fin.fromℕ≤″ _ lemma′) + div-helper 0 s d s * suc s ∎))
+(a divMod 0) {}
+(a divMod suc n) = result (a div suc n) (a mod suc n) $ begin
+ a ≡⟨ a≡a%n+[a/n]*n a n ⟩
+ a % suc n + [a/n]*n ≡⟨ cong (_+ [a/n]*n) (sym (toℕ-fromℕ≤ (a%n<n a n))) ⟩
+ toℕ (fromℕ≤ (a%n<n a n)) + [a/n]*n ∎
where
- lemma = s≤s (mod-lemma 0 d s)
- lemma′ = Nat.erase (≤⇒≤″ lemma)
+ [a/n]*n = a div suc n * suc n
diff --git a/src/Data/Nat/DivMod/Core.agda b/src/Data/Nat/DivMod/Core.agda
new file mode 100644
index 0000000..dfd6eb1
--- /dev/null
+++ b/src/Data/Nat/DivMod/Core.agda
@@ -0,0 +1,102 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Core lemmas for division and modulus operations
+------------------------------------------------------------------------
+
+module Data.Nat.DivMod.Core where
+
+open import Agda.Builtin.Nat using ()
+ renaming (div-helper to divₕ; mod-helper to modₕ)
+
+open import Data.Nat
+open import Data.Nat.Properties
+open import Relation.Binary.PropositionalEquality
+
+open ≡-Reasoning
+
+-------------------------------------------------------------------------
+-- mod lemmas
+
+modₕ-skipTo0 : ∀ acc n a b → modₕ acc n (b + a) a ≡ modₕ (a + acc) n b 0
+modₕ-skipTo0 acc n zero b = cong (λ v → modₕ acc n v 0) (+-identityʳ b)
+modₕ-skipTo0 acc n (suc a) b = begin
+ modₕ acc n (b + suc a) (suc a) ≡⟨ cong (λ v → modₕ acc n v (suc a)) (+-suc b a) ⟩
+ modₕ acc n (suc b + a) (suc a) ≡⟨⟩
+ modₕ (suc acc) n (b + a) a ≡⟨ modₕ-skipTo0 (suc acc) n a b ⟩
+ modₕ (a + suc acc) n b 0 ≡⟨ cong (λ v → modₕ v n b 0) (+-suc a acc) ⟩
+ modₕ (suc a + acc) n b 0 ∎
+
+modₕ-skipToEnd : ∀ acc n a b → modₕ acc n a (a + b) ≡ acc + a
+modₕ-skipToEnd acc n zero b = sym (+-identityʳ acc)
+modₕ-skipToEnd acc n (suc a) b = begin
+ modₕ (suc acc) n a (a + b) ≡⟨ modₕ-skipToEnd (suc acc) n a b ⟩
+ suc acc + a ≡⟨ sym (+-suc acc a) ⟩
+ acc + suc a ∎
+
+a[modₕ]1≡0 : ∀ a → modₕ 0 0 a 0 ≡ 0
+a[modₕ]1≡0 zero = refl
+a[modₕ]1≡0 (suc a) = a[modₕ]1≡0 a
+
+n[modₕ]n≡0 : ∀ acc v → modₕ acc (acc + v) (suc v) v ≡ 0
+n[modₕ]n≡0 acc v = modₕ-skipTo0 acc (acc + v) v 1
+
+a[modₕ]n<n : ∀ acc d n → modₕ acc (acc + n) d n ≤ acc + n
+a[modₕ]n<n acc zero n = m≤m+n acc n
+a[modₕ]n<n acc (suc d) zero = a[modₕ]n<n zero d (acc + 0)
+a[modₕ]n<n acc (suc d) (suc n)
+ rewrite +-suc acc n = a[modₕ]n<n (suc acc) d n
+
+modₕ-idem : ∀ acc a n → modₕ 0 (acc + n) (modₕ acc (acc + n) a n) (acc + n) ≡ modₕ acc (acc + n) a n
+modₕ-idem acc zero n = modₕ-skipToEnd 0 (acc + n) acc n
+modₕ-idem acc (suc a) zero rewrite +-identityʳ acc = modₕ-idem 0 a acc
+modₕ-idem acc (suc a) (suc n) rewrite +-suc acc n = modₕ-idem (suc acc) a n
+
+a+n[modₕ]n≡a[modₕ]n : ∀ acc a n → modₕ acc (acc + n) (acc + a + suc n) n ≡ modₕ acc (acc + n) a n
+a+n[modₕ]n≡a[modₕ]n acc zero n rewrite +-identityʳ acc = begin
+ modₕ acc (acc + n) (acc + suc n) n ≡⟨ cong (λ v → modₕ acc (acc + n) v n) (+-suc acc n) ⟩
+ modₕ acc (acc + n) (suc acc + n) n ≡⟨ modₕ-skipTo0 acc (acc + n) n (suc acc) ⟩
+ modₕ (acc + n) (acc + n) (suc acc) 0 ≡⟨⟩
+ modₕ 0 (acc + n) acc (acc + n) ≡⟨ modₕ-skipToEnd 0 (acc + n) acc n ⟩
+ acc ∎
+a+n[modₕ]n≡a[modₕ]n acc (suc a) zero rewrite +-identityʳ acc = begin
+ modₕ acc acc (acc + suc a + 1) 0 ≡⟨ cong (λ v → modₕ acc acc v 0) (+-comm (acc + suc a) 1) ⟩
+ modₕ acc acc (1 + (acc + suc a)) 0 ≡⟨⟩
+ modₕ 0 acc (acc + suc a) acc ≡⟨ cong (λ v → modₕ 0 acc v acc) (+-comm acc (suc a)) ⟩
+ modₕ 0 acc (suc a + acc) acc ≡⟨ cong (λ v → modₕ 0 acc v acc) (sym (+-suc a acc)) ⟩
+ modₕ 0 acc (a + suc acc) acc ≡⟨ a+n[modₕ]n≡a[modₕ]n 0 a acc ⟩
+ modₕ 0 acc a acc ∎
+a+n[modₕ]n≡a[modₕ]n acc (suc a) (suc n) rewrite +-suc acc n = begin
+ mod₁ (acc + suc a + (2 + n)) (suc n) ≡⟨ cong (λ v → mod₁ (v + suc (suc n)) (suc n)) (+-suc acc a) ⟩
+ mod₁ (suc acc + a + (2 + n)) (suc n) ≡⟨⟩
+ mod₂ (acc + a + (2 + n)) n ≡⟨ cong (λ v → mod₂ v n) (sym (+-assoc (acc + a) 1 (suc n))) ⟩
+ mod₂ (acc + a + 1 + suc n) n ≡⟨ cong (λ v → mod₂ (v + suc n) n) (+-comm (acc + a) 1) ⟩
+ mod₂ (suc acc + a + suc n) n ≡⟨ a+n[modₕ]n≡a[modₕ]n (suc acc) a n ⟩
+ mod₂ a n ∎
+ where
+ mod₁ = modₕ acc (suc acc + n)
+ mod₂ = modₕ (suc acc) (suc acc + n)
+
+-------------------------------------------------------------------------
+-- division lemmas
+
+-- The quotient and remainder are related to the dividend and
+-- divisor in the right way.
+
+division-lemma : ∀ accᵐ accᵈ d n →
+ accᵐ + accᵈ * suc (accᵐ + n) + d ≡
+ modₕ accᵐ (accᵐ + n) d n + divₕ accᵈ (accᵐ + n) d n * suc (accᵐ + n)
+division-lemma accᵐ accᵈ zero n = +-identityʳ _
+division-lemma accᵐ accᵈ (suc d) zero rewrite +-identityʳ accᵐ = begin
+ accᵐ + accᵈ * suc accᵐ + suc d ≡⟨ +-suc _ d ⟩
+ suc accᵈ * suc accᵐ + d ≡⟨ division-lemma zero (suc accᵈ) d accᵐ ⟩
+ modₕ 0 accᵐ d accᵐ +
+ divₕ (suc accᵈ) accᵐ d accᵐ * suc accᵐ ≡⟨⟩
+ modₕ accᵐ accᵐ (suc d) 0 +
+ divₕ accᵈ accᵐ (suc d) 0 * suc accᵐ ∎
+division-lemma accᵐ accᵈ (suc d) (suc n) rewrite +-suc accᵐ n = begin
+ accᵐ + accᵈ * m + suc d ≡⟨ +-suc _ d ⟩
+ suc (accᵐ + accᵈ * m + d) ≡⟨ division-lemma (suc accᵐ) accᵈ d n ⟩
+ modₕ _ _ d n + divₕ accᵈ _ d n * m ∎
+ where
+ m = 2 + accᵐ + n
diff --git a/src/Data/Nat/DivMod/Unsafe.agda b/src/Data/Nat/DivMod/Unsafe.agda
new file mode 100644
index 0000000..4e852d1
--- /dev/null
+++ b/src/Data/Nat/DivMod/Unsafe.agda
@@ -0,0 +1,45 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- More efficient (and unsafe) mod and divMod operations
+------------------------------------------------------------------------
+
+module Data.Nat.DivMod.Unsafe where
+
+open import Data.Nat using (ℕ; _+_; _*_; _≟_; zero; suc)
+open import Data.Nat.DivMod hiding (_mod_; _divMod_)
+open import Data.Nat.Properties using (≤⇒≤″)
+import Data.Nat.Unsafe as NatUnsafe
+open import Data.Fin using (Fin; toℕ; fromℕ≤″)
+open import Data.Fin.Properties using (toℕ-fromℕ≤″)
+open import Function using (_$_)
+open import Relation.Nullary.Decidable using (False)
+open import Relation.Binary.PropositionalEquality
+ using (refl; sym; cong; module ≡-Reasoning)
+import Relation.Binary.PropositionalEquality.TrustMe as TrustMe
+ using (erase)
+
+open ≡-Reasoning
+
+infixl 7 _mod_ _divMod_
+
+------------------------------------------------------------------------
+-- Certified modulus
+
+_mod_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} → Fin divisor
+(a mod 0) {}
+(a mod suc n) = fromℕ≤″ (a % suc n) (NatUnsafe.erase (≤⇒≤″ (a%n<n a n)))
+
+------------------------------------------------------------------------
+-- Returns modulus and division result with correctness proof
+
+_divMod_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} →
+ DivMod dividend divisor
+(a divMod 0) {}
+(a divMod suc n) = result (a div suc n) (a mod suc n) $ TrustMe.erase $ begin
+ a ≡⟨ a≡a%n+[a/n]*n a n ⟩
+ a % suc n + [a/n]*n ≡⟨ cong (_+ [a/n]*n) (sym (toℕ-fromℕ≤″ lemma′)) ⟩
+ toℕ (fromℕ≤″ _ lemma′) + [a/n]*n ∎
+ where
+ lemma′ = NatUnsafe.erase (≤⇒≤″ (a%n<n a n))
+ [a/n]*n = a div suc n * suc n
diff --git a/src/Data/Nat/Divisibility.agda b/src/Data/Nat/Divisibility.agda
index 8faaa60..ac6a544 100644
--- a/src/Data/Nat/Divisibility.agda
+++ b/src/Data/Nat/Divisibility.agda
@@ -6,20 +6,24 @@
module Data.Nat.Divisibility where
+open import Algebra
open import Data.Nat as Nat
open import Data.Nat.DivMod
open import Data.Nat.Properties
+open import Data.Nat.Solver
open import Data.Fin using (Fin; zero; suc; toℕ)
import Data.Fin.Properties as FP
-open SemiringSolver
-open import Algebra
open import Data.Product
+open import Function
+open import Function.Equivalence using (_⇔_; equivalence)
open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
open import Relation.Binary
import Relation.Binary.PartialOrderReasoning as PartialOrderReasoning
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; _≢_; refl; sym; trans; cong; cong₂; subst)
-open import Function
+
+open +-*-Solver
------------------------------------------------------------------------
-- m ∣ n is inhabited iff m divides n. Some sources, like Hardy and
@@ -30,8 +34,10 @@ open import Function
infix 4 _∣_ _∤_
-data _∣_ : ℕ → ℕ → Set where
- divides : {m n : ℕ} (q : ℕ) (eq : n ≡ q * m) → m ∣ n
+record _∣_ (m n : ℕ) : Set where
+ constructor divides
+ field quotient : ℕ
+ equality : n ≡ quotient * m
_∤_ : Rel ℕ _
m ∤ n = ¬ (m ∣ n)
@@ -61,11 +67,10 @@ quotient (divides q _) = q
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₂))
+∣-antisym {m} {0} _ (divides q eq) = trans eq (*-comm q 0)
+∣-antisym {0} {n} (divides p eq) _ = sym (trans eq (*-comm p 0))
+∣-antisym {suc m} {suc n} (divides p eq₁) (divides q eq₂) =
+ ≤-antisym (∣⇒≤ (divides p eq₁)) (∣⇒≤ (divides q eq₂))
∣-isPreorder : IsPreorder _≡_ _∣_
∣-isPreorder = record
@@ -89,7 +94,8 @@ poset = record
}
module ∣-Reasoning = PartialOrderReasoning poset
- renaming (_≤⟨_⟩_ to _∣⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
+ hiding (_≈⟨_⟩_)
+ renaming (_≤⟨_⟩_ to _∣⟨_⟩_)
------------------------------------------------------------------------
-- Simple properties of _∣_
@@ -105,8 +111,11 @@ n ∣0 = divides 0 refl
n∣n : ∀ {n} → n ∣ n
n∣n = ∣-refl
-n|m*n : ∀ m {n} → n ∣ m * n
-n|m*n m = divides m refl
+n∣m*n : ∀ m {n} → n ∣ m * n
+n∣m*n m = divides m refl
+
+m∣m*n : ∀ {m} n → m ∣ m * n
+m∣m*n n = divides n (*-comm _ n)
0∣⇒≡0 : ∀ {n} → 0 ∣ n → n ≡ 0
0∣⇒≡0 {n} 0∣n = ∣-antisym (n ∣0) 0∣n
@@ -117,92 +126,128 @@ n|m*n m = divides m 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
+module _ where
+
+ open PropEq.≡-Reasoning
+
+ ∣m⇒∣m*n : ∀ {i m} n → i ∣ m → i ∣ m * n
+ ∣m⇒∣m*n {i} {m} n (divides q eq) = divides (q * n) $ begin
+ m * n ≡⟨ cong (_* n) eq ⟩
+ q * i * n ≡⟨ *-assoc q i n ⟩
+ q * (i * n) ≡⟨ cong (q *_) (*-comm i n) ⟩
+ q * (n * i) ≡⟨ sym (*-assoc q n i) ⟩
+ q * n * i ∎
+
+ ∣n⇒∣m*n : ∀ {i} m {n} → i ∣ n → i ∣ m * n
+ ∣n⇒∣m*n {i} m {n} (divides q eq) = divides (m * q) $ begin
+ m * n ≡⟨ cong (m *_) eq ⟩
+ m * (q * i) ≡⟨ sym (*-assoc m q i) ⟩
+ m * q * i ∎
+
+ ∣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 ∎
+
+ ∣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 ∎
+
+ *-cong : ∀ {i j} k → i ∣ j → k * i ∣ k * 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) ∎
+
+ /-cong : ∀ {i j} k → suc k * i ∣ suc k * j → i ∣ j
+ /-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) ∎))
+
+ m%n≡0⇒n∣m : ∀ m n → m % (suc n) ≡ 0 → suc n ∣ m
+ m%n≡0⇒n∣m m n eq = divides (m div (suc n)) (begin
+ m ≡⟨ a≡a%n+[a/n]*n m n ⟩
+ m % (suc n) + m div (suc n) * (suc n) ≡⟨ cong₂ _+_ eq refl ⟩
+ m div (suc n) * (suc n) ∎)
+
+ n∣m⇒m%n≡0 : ∀ m n → suc n ∣ m → m % (suc n) ≡ 0
+ n∣m⇒m%n≡0 m n (divides v eq) = begin
+ m % (suc n) ≡⟨ cong (_% (suc n)) eq ⟩
+ (v * suc n) % (suc n) ≡⟨ kn%n≡0 v n ⟩
+ 0 ∎
+
+ m%n≡0⇔n∣m : ∀ m n → m % (suc n) ≡ 0 ⇔ suc n ∣ m
+ m%n≡0⇔n∣m m n = equivalence (m%n≡0⇒n∣m m n) (n∣m⇒m%n≡0 m n)
-∣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
+-- Divisibility is decidable.
-*-cong : ∀ {i j} k → i ∣ j → k * i ∣ k * 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
+infix 4 _∣?_
-/-cong : ∀ {i j} k → suc k * i ∣ suc k * j → i ∣ j
-/-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
+_∣?_ : Decidable _∣_
+zero ∣? zero = yes (0 ∣0)
+zero ∣? suc m = no ((λ()) ∘′ 0∣⇒≡0)
+suc n ∣? m = Dec.map (m%n≡0⇔n∣m m n) (m % (suc n) ≟ 0)
--- If the remainder after division is non-zero, then the divisor does
--- not divide the dividend.
+------------------------------------------------------------------------
+-- DEPRECATED - please use new names as continuing support for the old
+-- names is not guaranteed.
-nonZeroDivisor-lemma
- : ∀ m q (r : Fin (1 + m)) → toℕ r ≢ 0 →
- 1 + m ∤ toℕ r + q * (1 + m)
+∣-+ = ∣m∣n⇒∣m+n
+{-# WARNING_ON_USAGE ∣-+
+"Warning: ∣-+ was deprecated in v0.14.
+Please use ∣m∣n⇒∣m+n instead."
+#-}
+∣-∸ = ∣m+n∣m⇒∣n
+{-# WARNING_ON_USAGE ∣-∸
+"Warning: ∣-∸ was deprecated in v0.14.
+Please use ∣m+n∣m⇒∣n instead."
+#-}
+∣-* = n∣m*n
+{-# WARNING_ON_USAGE ∣-*
+"Warning: ∣-* was deprecated in v0.14.
+Please use n∣m*n instead."
+#-}
+
+nonZeroDivisor-lemma : ∀ 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
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) =
- ¬i+1+j≤i m $ begin
+ 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 ⟩
+ toℕ r ≤⟨ FP.toℕ≤pred[n] r ⟩
m ∎
where open ≤-Reasoning
nonZeroDivisor-lemma m (suc q) r r≢zero d =
- nonZeroDivisor-lemma m q r r≢zero (∣m+n|m⇒|n d' ∣-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) (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 =
- yes $ divides q 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
+{-# WARNING_ON_USAGE nonZeroDivisor-lemma
+"Warning: nonZeroDivisor-lemma was deprecated in v0.17."
+#-}
diff --git a/src/Data/Nat/GCD.agda b/src/Data/Nat/GCD.agda
index de37813..a8ad363 100644
--- a/src/Data/Nat/GCD.agda
+++ b/src/Data/Nat/GCD.agda
@@ -7,17 +7,17 @@
module Data.Nat.GCD where
open import Data.Nat
-open import Data.Nat.Divisibility as Div
-open import Relation.Binary
-private module P = Poset Div.poset
+open import Data.Nat.Divisibility
+open import Data.Nat.GCD.Lemmas
+open import Data.Nat.Properties using (+-suc)
open import Data.Product
-open import Relation.Binary.PropositionalEquality as PropEq using (_≡_; subst)
-open import Relation.Nullary using (Dec; yes; no)
+open import Function
open import Induction
-open import Induction.Nat using (<′-Rec; <′-rec-builder)
+open import Induction.Nat using (<′-Rec; <′-recBuilder)
open import Induction.Lexicographic
-open import Function
-open import Data.Nat.GCD.Lemmas
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_; subst)
+open import Relation.Nullary using (Dec; yes; no)
------------------------------------------------------------------------
-- Greatest common divisor
@@ -42,7 +42,7 @@ module GCD where
-- The gcd is unique.
unique : ∀ {d₁ d₂ m n} → GCD m n d₁ → GCD m n d₂ → d₁ ≡ d₂
- unique d₁ d₂ = P.antisym (GCD.greatest d₂ (GCD.commonDivisor d₁))
+ unique d₁ d₂ = ∣-antisym (GCD.greatest d₂ (GCD.commonDivisor d₁))
(GCD.greatest d₁ (GCD.commonDivisor d₂))
-- The gcd relation is "symmetric".
@@ -53,12 +53,12 @@ module GCD where
-- The gcd relation is "reflexive".
refl : ∀ {n} → GCD n n n
- refl = is (P.refl , P.refl) proj₁
+ refl = is (∣-refl , ∣-refl) proj₁
-- The GCD of 0 and n is n.
base : ∀ {n} → GCD 0 n n
- base {n} = is (n ∣0 , P.refl) proj₂
+ base {n} = is (n ∣0 , ∣-refl) proj₂
-- If d is the gcd of n and k, then it is also the gcd of n and
-- n + k.
@@ -68,7 +68,7 @@ module GCD where
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₁ , ∣m+n|m⇒|n d₂ d₁)
+ greatest′ (d₁ , d₂) = GCD.greatest g (d₁ , ∣m+n∣m⇒∣n d₂ d₁)
open GCD public using (GCD) hiding (module GCD)
@@ -99,10 +99,10 @@ module Bézout where
sym (-+ x y eq) = +- y x eq
refl : ∀ {d} → Identity d d d
- refl = -+ 0 1 PropEq.refl
+ refl = -+ 0 1 P.refl
base : ∀ {d} → Identity d 0 d
- base = -+ 0 1 PropEq.refl
+ base = -+ 0 1 P.refl
private
infixl 7 _⊕_
@@ -143,7 +143,7 @@ module Bézout where
stepˡ : ∀ {n k} → Lemma n (suc k) → Lemma n (suc (n + k))
stepˡ {n} {k} (result d g b) =
- PropEq.subst (Lemma n) (lem₀ n k) $
+ subst (Lemma n) (+-suc n k) $
result d (GCD.step g) (Identity.step b)
stepʳ : ∀ {n k} → Lemma (suc k) n → Lemma (suc (n + k)) n
@@ -155,7 +155,7 @@ 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 [ <′-recBuilder ⊗ <′-recBuilder ] P gcd (m , n)
where
P : ℕ × ℕ → Set
P (m , n) = Lemma m n
@@ -176,19 +176,19 @@ module Bézout where
identity : ∀ {m n d} → GCD m n d → Identity d m n
identity {m} {n} g with lemma m n
- identity g | result d g′ b with GCD.unique g g′
- identity g | result d g′ b | PropEq.refl = b
+ ... | result d g′ b with GCD.unique g g′
+ ... | P.refl = b
-- Calculates the gcd of the arguments.
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)
+... | 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)
+... | no ¬g = no (¬g ∘ GCD.unique 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 188e7d4..b0f3ff2 100644
--- a/src/Data/Nat/GCD/Lemmas.agda
+++ b/src/Data/Nat/GCD/Lemmas.agda
@@ -7,23 +7,41 @@
module Data.Nat.GCD.Lemmas where
open import Data.Nat
-import Data.Nat.Properties as NatProp
-open NatProp.SemiringSolver
+open import Data.Nat.Properties
+open import Data.Nat.Solver
+open import Function
open import Relation.Binary.PropositionalEquality
+
+open +-*-Solver
open ≡-Reasoning
-open import Function
-lem₀ = solve 2 (λ n k → n :+ (con 1 :+ k) := con 1 :+ n :+ k) refl
+private
+ distrib-comm : ∀ x k n → x * k + x * n ≡ x * (n + k)
+ distrib-comm =
+ solve 3 (λ x k n → x :* k :+ x :* n := x :* (n :+ k)) refl
+
+ distrib-comm₂ : ∀ d x k n → d + x * (n + k) ≡ d + x * k + x * n
+ distrib-comm₂ =
+ solve 4 (λ d x k n → d :+ x :* (n :+ k) := d :+ x :* k :+ x :* n) refl
+
+-- Other properties
+-- TODO: Can this proof be simplified? An automatic solver which can
+-- handle ∸ would be nice...
+lem₀ : ∀ i j m n → i * m ≡ j * m + n → (i ∸ j) * m ≡ n
+lem₀ i j m n eq = begin
+ (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 ∎
lem₁ : ∀ i j → 2 + i ≤′ 2 + j + i
-lem₁ i j = NatProp.≤⇒≤′ $ s≤s $ s≤s $ NatProp.n≤m+n j i
+lem₁ i j = ≤⇒≤′ $ s≤s $ s≤s $ n≤m+n j i
lem₂ : ∀ d x {k n} →
d + x * k ≡ x * n → d + x * (n + k) ≡ 2 * x * n
lem₂ d x {k} {n} eq = begin
- d + x * (n + k) ≡⟨ solve 4 (λ d x n k → d :+ x :* (n :+ k)
- := d :+ x :* k :+ x :* n)
- refl d x n k ⟩
+ d + x * (n + k) ≡⟨ distrib-comm₂ d x k n ⟩
d + x * k + x * n ≡⟨ cong₂ _+_ eq refl ⟩
x * n + x * n ≡⟨ solve 3 (λ x n k → x :* n :+ x :* n
:= con 2 :* x :* n)
@@ -34,10 +52,8 @@ lem₃ : ∀ d x {i k n} →
d + (1 + x + i) * k ≡ x * n →
d + (1 + x + i) * (n + k) ≡ (1 + 2 * x + i) * n
lem₃ d x {i} {k} {n} eq = begin
- d + y * (n + k) ≡⟨ solve 4 (λ d y n k → d :+ y :* (n :+ k)
- := (d :+ y :* k) :+ y :* n)
- refl d y n k ⟩
- (d + y * k) + y * n ≡⟨ cong₂ _+_ eq refl ⟩
+ d + y * (n + k) ≡⟨ distrib-comm₂ d y k n ⟩
+ d + y * k + y * n ≡⟨ cong₂ _+_ eq refl ⟩
x * n + y * n ≡⟨ solve 3 (λ x n i → x :* n :+ (con 1 :+ x :+ i) :* n
:= (con 1 :+ con 2 :* x :+ i) :* n)
refl x n i ⟩
@@ -48,19 +64,13 @@ lem₄ : ∀ d y {k i} n →
d + y * k ≡ (1 + y + i) * n →
d + y * (n + k) ≡ (1 + 2 * y + i) * n
lem₄ d y {k} {i} n eq = begin
- d + y * (n + k) ≡⟨ solve 4 (λ d y n k → d :+ y :* (n :+ k)
- := d :+ y :* k :+ y :* n)
- refl d y n k ⟩
+ d + y * (n + k) ≡⟨ distrib-comm₂ d y k n ⟩
d + y * k + y * n ≡⟨ cong₂ _+_ eq refl ⟩
(1 + y + i) * n + y * n ≡⟨ solve 3 (λ y i n → (con 1 :+ y :+ i) :* n :+ y :* n
:= (con 1 :+ con 2 :* y :+ i) :* n)
refl y i n ⟩
(1 + 2 * y + i) * n ∎
-private
- distrib-comm =
- solve 3 (λ x k n → x :* k :+ x :* n := x :* (n :+ k)) refl
-
lem₅ : ∀ d x {n k} →
d + x * n ≡ x * k →
d + 2 * x * n ≡ x * (n + k)
@@ -99,13 +109,13 @@ lem₈ : ∀ {i j k q} x y →
1 + y * j ≡ x * i → j * k ≡ q * i →
k ≡ (x * k ∸ y * q) * i
lem₈ {i} {j} {k} {q} x y eq eq′ =
- sym (NatProp.im≡jm+n⇒[i∸j]m≡n (x * k) (y * q) i k lemma)
+ sym (lem₀ (x * k) (y * q) i k lemma)
where
lemma = begin
x * k * i ≡⟨ solve 3 (λ x k i → x :* k :* i
:= x :* i :* k)
refl x k i ⟩
- x * i * k ≡⟨ cong (λ n → n * k) (sym eq) ⟩
+ x * i * k ≡⟨ cong (_* k) (sym eq) ⟩
(1 + y * j) * k ≡⟨ solve 3 (λ y j k → (con 1 :+ y :* j) :* k
:= y :* (j :* k) :+ k)
refl y j k ⟩
@@ -119,7 +129,7 @@ lem₉ : ∀ {i j k q} x y →
1 + x * i ≡ y * j → j * k ≡ q * i →
k ≡ (y * q ∸ x * k) * i
lem₉ {i} {j} {k} {q} x y eq eq′ =
- sym (NatProp.im≡jm+n⇒[i∸j]m≡n (y * q) (x * k) i k lemma)
+ sym (lem₀ (y * q) (x * k) i k lemma)
where
lem = solve 3 (λ a b c → a :* b :* c := b :* c :* a) refl
lemma = begin
@@ -136,9 +146,9 @@ lem₁₀ : ∀ {a′} b c {d} e f → let a = suc a′ in
a + b * (c * d * a) ≡ e * (f * d * a) →
d ≡ 1
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ʳ-≡ (e * f * d) (b * c * d + 1) (begin
+ i*j≡1⇒j≡1 (e * f ∸ b * c) d
+ (lem₀ (e * f) (b * c) d 1
+ (*-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 ⟩
@@ -153,18 +163,14 @@ lem₁₁ : ∀ {i j m n k d} x y →
1 + y * j ≡ x * i → i * k ≡ m * d → j * k ≡ n * d →
k ≡ (x * m ∸ y * n) * d
lem₁₁ {i} {j} {m} {n} {k} {d} x y eq eq₁ eq₂ =
- sym (NatProp.im≡jm+n⇒[i∸j]m≡n (x * m) (y * n) d k lemma)
- where
- assoc = solve 3 (λ x y z → x :* y :* z := x :* (y :* z)) refl
-
- lemma = begin
- x * m * d ≡⟨ assoc x m d ⟩
- x * (m * d) ≡⟨ cong (_*_ x) (sym eq₁) ⟩
- x * (i * k) ≡⟨ sym (assoc x i k) ⟩
+ sym (lem₀ (x * m) (y * n) d k (begin
+ x * m * d ≡⟨ *-assoc x m d ⟩
+ x * (m * d) ≡⟨ cong (x *_) (sym eq₁) ⟩
+ x * (i * k) ≡⟨ sym (*-assoc x i k) ⟩
x * i * k ≡⟨ cong₂ _*_ (sym eq) refl ⟩
(1 + y * j) * k ≡⟨ solve 3 (λ y j k → (con 1 :+ y :* j) :* k
:= y :* (j :* k) :+ k)
refl y j k ⟩
y * (j * k) + k ≡⟨ cong (λ p → y * p + k) eq₂ ⟩
- y * (n * d) + k ≡⟨ cong₂ _+_ (sym $ assoc y n d) refl ⟩
- y * n * d + k ∎
+ y * (n * d) + k ≡⟨ cong₂ _+_ (sym $ *-assoc y n d) refl ⟩
+ y * n * d + k ∎))
diff --git a/src/Data/Nat/InfinitelyOften.agda b/src/Data/Nat/InfinitelyOften.agda
index dfeb6c7..f3e6883 100644
--- a/src/Data/Nat/InfinitelyOften.agda
+++ b/src/Data/Nat/InfinitelyOften.agda
@@ -6,30 +6,34 @@
module Data.Nat.InfinitelyOften where
-import Level
-open import Algebra
-open import Category.Monad
-open import Data.Empty
-open import Function
+open import Category.Monad using (RawMonad)
+open import Level using (0ℓ)
+open import Data.Empty using (⊥-elim)
open import Data.Nat
open import Data.Nat.Properties
open import Data.Product as Prod hiding (map)
open import Data.Sum hiding (map)
+open import Function
open import Relation.Binary.PropositionalEquality
-open import Relation.Nullary
-open import Relation.Nullary.Negation
-open import Relation.Unary using (_∪_; _⊆_)
-open RawMonad (¬¬-Monad {p = Level.zero})
+open import Relation.Nullary using (¬_)
+open import Relation.Nullary.Negation using (¬¬-Monad; call/cc)
+open import Relation.Unary using (Pred; _∪_; _⊆_)
+open RawMonad (¬¬-Monad {p = 0ℓ})
-- Only true finitely often.
-Fin : (ℕ → Set) → Set
+Fin : ∀ {ℓ} → Pred ℕ ℓ → Set ℓ
Fin P = ∃ λ i → ∀ j → i ≤ j → ¬ P j
+-- A non-constructive definition of "true infinitely often".
+
+Inf : ∀ {ℓ} → Pred ℕ ℓ → Set ℓ
+Inf P = ¬ Fin P
+
-- Fin is preserved by binary sums.
-_∪-Fin_ : ∀ {P Q} → Fin P → Fin Q → Fin (P ∪ Q)
-_∪-Fin_ {P} {Q} (i , ¬p) (j , ¬q) = (i ⊔ j , helper)
+_∪-Fin_ : ∀ {ℓp ℓq P Q} → Fin {ℓp} P → Fin {ℓq} Q → Fin (P ∪ Q)
+_∪-Fin_ {P = P} {Q} (i , ¬p) (j , ¬q) = (i ⊔ j , helper)
where
open ≤-Reasoning
@@ -44,11 +48,6 @@ _∪-Fin_ {P} {Q} (i , ¬p) (j , ¬q) = (i ⊔ j , helper)
i ⊔ j ≤⟨ i⊔j≤k ⟩
k ∎) q
--- A non-constructive definition of "true infinitely often".
-
-Inf : (ℕ → Set) → Set
-Inf P = ¬ Fin P
-
-- Inf commutes with binary sums (in the double-negation monad).
commutes-with-∪ : ∀ {P Q} → Inf (P ∪ Q) → ¬ ¬ (Inf P ⊎ Inf Q)
@@ -59,14 +58,14 @@ commutes-with-∪ p∪q =
-- Inf is functorial.
-map : ∀ {P Q} → P ⊆ Q → Inf P → Inf Q
+map : ∀ {ℓp ℓq P Q} → P ⊆ Q → Inf {ℓp} P → Inf {ℓq} Q
map P⊆Q ¬fin = ¬fin ∘ Prod.map id (λ fin j i≤j → fin j i≤j ∘ P⊆Q)
-- Inf is upwards closed.
-up : ∀ {P} n → Inf P → Inf (P ∘ _+_ n)
+up : ∀ {ℓ P} n → Inf {ℓ} P → Inf (P ∘ _+_ n)
up zero = id
-up {P} (suc n) = up n ∘ up₁
+up {P = P} (suc n) = up n ∘ up₁
where
up₁ : Inf P → Inf (P ∘ suc)
up₁ ¬fin (i , fin) = ¬fin (suc i , helper)
@@ -76,7 +75,7 @@ up {P} (suc n) = up n ∘ up₁
-- A witness.
-witness : ∀ {P} → Inf P → ¬ ¬ ∃ P
+witness : ∀ {ℓ P} → Inf {ℓ} P → ¬ ¬ ∃ P
witness ¬fin ¬p = ¬fin (0 , λ i _ Pi → ¬p (i , Pi))
-- Two different witnesses.
diff --git a/src/Data/Nat/LCM.agda b/src/Data/Nat/LCM.agda
index e044d33..e0688c6 100644
--- a/src/Data/Nat/LCM.agda
+++ b/src/Data/Nat/LCM.agda
@@ -6,9 +6,10 @@
module Data.Nat.LCM where
+open import Algebra
open import Data.Nat
-import Data.Nat.Properties as NatProp
-open NatProp.SemiringSolver
+open import Data.Nat.Properties
+open import Data.Nat.Solver
open import Data.Nat.GCD
open import Data.Nat.Divisibility as Div
open import Data.Nat.Coprimality as Coprime
@@ -16,8 +17,10 @@ open import Data.Product
open import Function
open import Relation.Binary.PropositionalEquality as PropEq
using (_≡_; refl)
-open import Algebra
open import Relation.Binary
+
+open +-*-Solver
+
private
module P = Poset Div.poset
@@ -90,7 +93,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ʳ-≡ _ _ (begin
+ (divides q₄ $ *-cancelʳ-≡ _ _ (begin
q₁ * q₃ * d′ ≡⟨ lem₁ q₁ q₃ d′ ⟩
q₃ * (q₁ * d′) ≡⟨ PropEq.sym eq₃ ⟩
m ≡⟨ eq₄ ⟩
diff --git a/src/Data/Nat/Literals.agda b/src/Data/Nat/Literals.agda
new file mode 100644
index 0000000..b084a3c
--- /dev/null
+++ b/src/Data/Nat/Literals.agda
@@ -0,0 +1,17 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Natural Literals
+------------------------------------------------------------------------
+
+module Data.Nat.Literals where
+
+open import Agda.Builtin.FromNat
+open import Agda.Builtin.Nat
+open import Data.Unit
+
+number : Number Nat
+number = record
+ { Constraint = λ _ → ⊤
+ ; fromNat = λ n → n
+ }
diff --git a/src/Data/Nat/Primality.agda b/src/Data/Nat/Primality.agda
index b632f9e..4ab2fa6 100644
--- a/src/Data/Nat/Primality.agda
+++ b/src/Data/Nat/Primality.agda
@@ -6,29 +6,29 @@
module Data.Nat.Primality where
-open import Data.Empty
-open import Data.Fin as Fin hiding (_+_)
-open import Data.Fin.Dec
-open import Data.Nat
-open import Data.Nat.Divisibility
-open import Relation.Nullary
-open import Relation.Nullary.Decidable
-open import Relation.Nullary.Negation
-open import Relation.Unary
+open import Data.Empty using (⊥)
+open import Data.Fin using (Fin; toℕ)
+open import Data.Fin.Properties using (all?)
+open import Data.Nat using (ℕ; suc; _+_)
+open import Data.Nat.Divisibility using (_∤_; _∣?_)
+open import Relation.Nullary using (yes; no)
+open import Relation.Nullary.Decidable using (from-yes)
+open import Relation.Nullary.Negation using (¬?)
+open import Relation.Unary using (Decidable)
-- Definition of primality.
Prime : ℕ → Set
Prime 0 = ⊥
Prime 1 = ⊥
-Prime (suc (suc n)) = (i : Fin n) → ¬ (2 + Fin.toℕ i ∣ 2 + n)
+Prime (suc (suc n)) = (i : Fin n) → 2 + toℕ i ∤ 2 + n
-- Decision procedure for primality.
prime? : Decidable Prime
prime? 0 = no λ()
prime? 1 = no λ()
-prime? (suc (suc n)) = all? λ _ → ¬? (_ ∣? _)
+prime? (suc (suc n)) = all? (λ _ → ¬? (_ ∣? _))
private
diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda
index f3df5b3..9e9ebf3 100644
--- a/src/Data/Nat/Properties.agda
+++ b/src/Data/Nat/Properties.agda
@@ -9,24 +9,26 @@
module Data.Nat.Properties where
-open import Relation.Binary
-open import Function
open import Algebra
-import Algebra.RingSolver.Simple as Solver
-import Algebra.RingSolver.AlmostCommutativeRing as ACR
-open import Algebra.Structures
-open import Data.Nat as Nat
+open import Algebra.Morphism
+open import Function
+open import Function.Injection using (_↣_)
+open import Data.Nat.Base
open import Data.Product
open import Data.Sum
+open import Level using (0ℓ)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality
open import Relation.Nullary
+open import Relation.Nullary.Decidable using (via-injection; map′)
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 import Algebra.Structures (_≡_ {A = ℕ})
open ≡-Reasoning
------------------------------------------------------------------------
@@ -35,13 +37,23 @@ open ≡-Reasoning
suc-injective : ∀ {m n} → suc m ≡ suc n → m ≡ n
suc-injective refl = refl
+infix 4 _≟_
+
+_≟_ : Decidable {A = ℕ} _≡_
+zero ≟ zero = yes refl
+zero ≟ suc n = no λ()
+suc m ≟ zero = no λ()
+suc m ≟ suc n with m ≟ n
+... | yes refl = yes refl
+... | no m≢n = no (m≢n ∘ suc-injective)
+
≡-isDecEquivalence : IsDecEquivalence (_≡_ {A = ℕ})
≡-isDecEquivalence = record
{ isEquivalence = isEquivalence
; _≟_ = _≟_
}
-≡-decSetoid : DecSetoid _ _
+≡-decSetoid : DecSetoid 0ℓ 0ℓ
≡-decSetoid = record
{ Carrier = ℕ
; _≈_ = _≡_
@@ -51,6 +63,9 @@ suc-injective refl = refl
------------------------------------------------------------------------
-- Properties of _≤_
+≤-pred : ∀ {m n} → suc m ≤ suc n → m ≤ n
+≤-pred (s≤s m≤n) = m≤n
+
-- Relation-theoretic properties of _≤_
≤-reflexive : _≡_ ⇒ _≤_
≤-reflexive {zero} refl = z≤n
@@ -75,6 +90,18 @@ suc-injective refl = refl
... | inj₁ m≤n = inj₁ (s≤s m≤n)
... | inj₂ n≤m = inj₂ (s≤s n≤m)
+infix 4 _≤?_ _≥?_
+
+_≤?_ : Decidable _≤_
+zero ≤? _ = yes z≤n
+suc m ≤? zero = no λ()
+suc m ≤? suc n with m ≤? n
+... | yes m≤n = yes (s≤s m≤n)
+... | no m≰n = no (m≰n ∘ ≤-pred)
+
+_≥?_ : Decidable _≥_
+_≥?_ = flip _≤?_
+
≤-isPreorder : IsPreorder _≡_ _≤_
≤-isPreorder = record
{ isEquivalence = isEquivalence
@@ -82,7 +109,7 @@ suc-injective refl = refl
; trans = ≤-trans
}
-≤-preorder : Preorder _ _ _
+≤-preorder : Preorder 0ℓ 0ℓ 0ℓ
≤-preorder = record
{ isPreorder = ≤-isPreorder
}
@@ -93,13 +120,18 @@ suc-injective refl = refl
; antisym = ≤-antisym
}
+≤-poset : Poset 0ℓ 0ℓ 0ℓ
+≤-poset = record
+ { isPartialOrder = ≤-isPartialOrder
+ }
+
≤-isTotalOrder : IsTotalOrder _≡_ _≤_
≤-isTotalOrder = record
{ isPartialOrder = ≤-isPartialOrder
; total = ≤-total
}
-≤-totalOrder : TotalOrder _ _ _
+≤-totalOrder : TotalOrder 0ℓ 0ℓ 0ℓ
≤-totalOrder = record
{ isTotalOrder = ≤-isTotalOrder
}
@@ -111,12 +143,19 @@ suc-injective refl = refl
; _≤?_ = _≤?_
}
-≤-decTotalOrder : DecTotalOrder _ _ _
+≤-decTotalOrder : DecTotalOrder 0ℓ 0ℓ 0ℓ
≤-decTotalOrder = record
{ isDecTotalOrder = ≤-isDecTotalOrder
}
-- Other properties of _≤_
+s≤s-injective : ∀ {m n} {p q : m ≤ n} → s≤s p ≡ s≤s q → p ≡ q
+s≤s-injective refl = refl
+
+≤-irrelevance : Irrelevant _≤_
+≤-irrelevance z≤n z≤n = refl
+≤-irrelevance (s≤s m≤n₁) (s≤s m≤n₂) = cong s≤s (≤-irrelevance m≤n₁ m≤n₂)
+
≤-step : ∀ {m n} → m ≤ n → m ≤ 1 + n
≤-step z≤n = z≤n
≤-step (s≤s m≤n) = s≤s (≤-step m≤n)
@@ -124,9 +163,12 @@ suc-injective refl = refl
n≤1+n : ∀ n → n ≤ 1 + n
n≤1+n _ = ≤-step ≤-refl
-1+n≰n : ∀ {n} → ¬ 1 + n ≤ n
+1+n≰n : ∀ {n} → 1 + n ≰ n
1+n≰n (s≤s le) = 1+n≰n le
+n≤0⇒n≡0 : ∀ {n} → n ≤ 0 → n ≡ 0
+n≤0⇒n≡0 z≤n = refl
+
pred-mono : pred Preserves _≤_ ⟶ _≤_
pred-mono z≤n = z≤n
pred-mono (s≤s le) = le
@@ -143,8 +185,6 @@ pred-mono (s≤s le) = 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
@@ -170,6 +210,30 @@ x <? y = suc x ≤? y
... | tri≈ ≰ ≡ ≱ = tri≈ (≰ ∘ ≤-pred) (cong suc ≡) (≱ ∘ ≤-pred)
... | tri> ≰ ≢ ≥ = tri> (≰ ∘ ≤-pred) (≢ ∘ suc-injective) (s≤s ≥)
+infix 4 _<?_ _>?_
+
+_<?_ : Decidable _<_
+x <? y = suc x ≤? y
+
+_>?_ : Decidable _>_
+_>?_ = flip _<?_
+
+<-resp₂-≡ : _<_ Respects₂ _≡_
+<-resp₂-≡ = subst (_ <_) , subst (_< _)
+
+<-isStrictPartialOrder : IsStrictPartialOrder _≡_ _<_
+<-isStrictPartialOrder = record
+ { isEquivalence = isEquivalence
+ ; irrefl = <-irrefl
+ ; trans = <-trans
+ ; <-resp-≈ = <-resp₂-≡
+ }
+
+<-strictPartialOrder : StrictPartialOrder 0ℓ 0ℓ 0ℓ
+<-strictPartialOrder = record
+ { isStrictPartialOrder = <-isStrictPartialOrder
+ }
+
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
<-isStrictTotalOrder = record
{ isEquivalence = isEquivalence
@@ -177,12 +241,15 @@ x <? y = suc x ≤? y
; compare = <-cmp
}
-<-strictTotalOrder : StrictTotalOrder _ _ _
+<-strictTotalOrder : StrictTotalOrder 0ℓ 0ℓ 0ℓ
<-strictTotalOrder = record
{ isStrictTotalOrder = <-isStrictTotalOrder
}
-- Other properties of _<_
+<-irrelevance : Irrelevant _<_
+<-irrelevance = ≤-irrelevance
+
<⇒≤pred : ∀ {m n} → m < n → m ≤ pred n
<⇒≤pred (s≤s le) = le
@@ -192,6 +259,10 @@ x <? y = suc x ≤? y
<⇒≢ : _<_ ⇒ _≢_
<⇒≢ m<n refl = 1+n≰n m<n
+≤⇒≯ : _≤_ ⇒ _≯_
+≤⇒≯ z≤n ()
+≤⇒≯ (s≤s m≤n) (s≤s n≤m) = ≤⇒≯ m≤n n≤m
+
<⇒≱ : _<_ ⇒ _≱_
<⇒≱ (s≤s m+1≤n) (s≤s n≤m) = <⇒≱ m+1≤n n≤m
@@ -214,11 +285,17 @@ x <? y = suc x ≤? y
≮⇒≥ {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))
+≤∧≢⇒< : ∀ {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))
+
+n≮n : ∀ n → n ≮ n
+n≮n n = <-irrefl (refl {x = n})
+
+m<n⇒n≢0 : ∀ {m n} → m < n → n ≢ 0
+m<n⇒n≢0 (s≤s m≤n) ()
------------------------------------------------------------------------
-- Properties of _≤′_
@@ -239,6 +316,25 @@ s≤′s (≤′-step m≤′n) = ≤′-step (s≤′s m≤′n)
≤⇒≤′ z≤n = z≤′n
≤⇒≤′ (s≤s m≤n) = s≤′s (≤⇒≤′ m≤n)
+≤′-step-injective : ∀ {m n} {p q : m ≤′ n} → ≤′-step p ≡ ≤′-step q → p ≡ q
+≤′-step-injective refl = refl
+
+-- Decidablity for _≤'_
+
+infix 4 _≤′?_ _<′?_ _≥′?_ _>′?_
+
+_≤′?_ : Decidable _≤′_
+x ≤′? y = map′ ≤⇒≤′ ≤′⇒≤ (x ≤? y)
+
+_<′?_ : Decidable _<′_
+x <′? y = suc x ≤′? y
+
+_≥′?_ : Decidable _≥′_
+_≥′?_ = flip _≤′?_
+
+_>′?_ : Decidable _>′_
+_>′?_ = flip _<′?_
+
------------------------------------------------------------------------
-- Properties of _≤″_
@@ -259,6 +355,22 @@ s≤′s (≤′-step m≤′n) = ≤′-step (s≤′s m≤′n)
proof z≤n = refl
proof (s≤s m≤n) = cong suc (proof m≤n)
+-- Decidablity for _≤″_
+
+infix 4 _≤″?_ _<″?_ _≥″?_ _>″?_
+
+_≤″?_ : Decidable _≤″_
+x ≤″? y = map′ ≤⇒≤″ ≤″⇒≤ (x ≤? y)
+
+_<″?_ : Decidable _<″_
+x <″? y = suc x ≤″? y
+
+_≥″?_ : Decidable _≥″_
+_≥″?_ = flip _≤″?_
+
+_>″?_ : Decidable _>″_
+_>″?_ = flip _<″?_
+
------------------------------------------------------------------------
-- Properties of _+_
@@ -289,20 +401,41 @@ s≤′s (≤′-step m≤′n) = ≤′-step (s≤′s m≤′n)
suc (n + m) ≡⟨ sym (+-suc n m) ⟩
n + suc m ∎
-+-isSemigroup : IsSemigroup _≡_ _+_
++-isSemigroup : IsSemigroup _+_
+-isSemigroup = record
{ isEquivalence = isEquivalence
; assoc = +-assoc
; ∙-cong = cong₂ _+_
}
-+-0-isCommutativeMonoid : IsCommutativeMonoid _≡_ _+_ 0
++-semigroup : Semigroup 0ℓ 0ℓ
++-semigroup = record
+ { isSemigroup = +-isSemigroup
+ }
+
++-0-isMonoid : IsMonoid _+_ 0
++-0-isMonoid = record
+ { isSemigroup = +-isSemigroup
+ ; identity = +-identity
+ }
+
++-0-monoid : Monoid 0ℓ 0ℓ
++-0-monoid = record
+ { isMonoid = +-0-isMonoid
+ }
+
++-0-isCommutativeMonoid : IsCommutativeMonoid _+_ 0
+-0-isCommutativeMonoid = record
{ isSemigroup = +-isSemigroup
; identityˡ = +-identityˡ
; comm = +-comm
}
++-0-commutativeMonoid : CommutativeMonoid 0ℓ 0ℓ
++-0-commutativeMonoid = record
+ { isCommutativeMonoid = +-0-isCommutativeMonoid
+ }
+
-- Other properties of _+_ and _≡_
+-cancelˡ-≡ : LeftCancellative _≡_ _+_
@@ -343,6 +476,13 @@ i+j≡0⇒j≡0 i {j} i+j≡0 = i+j≡0⇒i≡0 j (trans (+-comm j i) (i+j≡0))
+-cancel-≤ : Cancellative _≤_ _+_
+-cancel-≤ = +-cancelˡ-≤ , +-cancelʳ-≤
+≤-stepsˡ : ∀ {m n} o → m ≤ n → m ≤ o + n
+≤-stepsˡ zero m≤n = m≤n
+≤-stepsˡ (suc o) m≤n = ≤-step (≤-stepsˡ o m≤n)
+
+≤-stepsʳ : ∀ {m n} o → m ≤ n → m ≤ n + o
+≤-stepsʳ {m} o m≤n = subst (m ≤_) (+-comm o _) (≤-stepsˡ o m≤n)
+
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)
@@ -351,10 +491,6 @@ 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)
@@ -367,25 +503,41 @@ m+n≤o⇒n≤o (suc m) m+n<o = m+n≤o⇒n≤o m (<⇒≤ m+n<o)
+-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ˡ-≤ : ∀ n → (_+ n) Preserves _≤_ ⟶ _≤_
++-monoˡ-≤ n m≤o = +-mono-≤ m≤o (≤-refl {n})
+
++-monoʳ-≤ : ∀ n → (n +_) Preserves _≤_ ⟶ _≤_
++-monoʳ-≤ n m≤o = +-mono-≤ (≤-refl {n}) m≤o
-+-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-<-≤ {_} {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)
++-mono-< x≤y = +-mono-≤-< (<⇒≤ x≤y)
+
++-monoˡ-< : ∀ n → (_+ n) Preserves _<_ ⟶ _<_
++-monoˡ-< n = +-monoˡ-≤ n
-¬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)
++-monoʳ-< : ∀ n → (n +_) Preserves _<_ ⟶ _<_
++-monoʳ-< zero m≤o = m≤o
++-monoʳ-< (suc n) m≤o = s≤s (+-monoʳ-< n m≤o)
+
+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 zero n = n≮n n
m+n≮n (suc m) (suc n) (s≤s m+n<n) = m+n≮n m (suc n) (≤-step m+n<n)
+m+n≮m : ∀ m n → m + n ≮ m
+m+n≮m m n = subst (_≮ m) (+-comm n m) (m+n≮n n m)
+
m≤′m+n : ∀ m n → m ≤′ m + n
m≤′m+n m n = ≤⇒≤′ (m≤m+n m n)
@@ -460,21 +612,56 @@ n≤′m+n (suc m) n = ≤′-step (n≤′m+n m n)
n * o + m * (n * o) ≡⟨⟩
suc m * (n * o) ∎
-*-isSemigroup : IsSemigroup _≡_ _*_
+*-isSemigroup : IsSemigroup _*_
*-isSemigroup = record
{ isEquivalence = isEquivalence
; assoc = *-assoc
; ∙-cong = cong₂ _*_
}
-*-1-isCommutativeMonoid : IsCommutativeMonoid _≡_ _*_ 1
+*-semigroup : Semigroup 0ℓ 0ℓ
+*-semigroup = record
+ { isSemigroup = *-isSemigroup
+ }
+
+*-1-isMonoid : IsMonoid _*_ 1
+*-1-isMonoid = record
+ { isSemigroup = *-isSemigroup
+ ; identity = *-identity
+ }
+
+*-1-monoid : Monoid 0ℓ 0ℓ
+*-1-monoid = record
+ { isMonoid = *-1-isMonoid
+ }
+
+*-1-isCommutativeMonoid : IsCommutativeMonoid _*_ 1
*-1-isCommutativeMonoid = record
{ isSemigroup = *-isSemigroup
; identityˡ = *-identityˡ
; comm = *-comm
}
-*-+-isCommutativeSemiring : IsCommutativeSemiring _≡_ _+_ _*_ 0 1
+*-1-commutativeMonoid : CommutativeMonoid 0ℓ 0ℓ
+*-1-commutativeMonoid = record
+ { isCommutativeMonoid = *-1-isCommutativeMonoid
+ }
+
+*-+-isSemiring : IsSemiring _+_ _*_ 0 1
+*-+-isSemiring = record
+ { isSemiringWithoutAnnihilatingZero = record
+ { +-isCommutativeMonoid = +-0-isCommutativeMonoid
+ ; *-isMonoid = *-1-isMonoid
+ ; distrib = *-distrib-+ }
+ ; zero = *-zero
+ }
+
+*-+-semiring : Semiring 0ℓ 0ℓ
+*-+-semiring = record
+ { isSemiring = *-+-isSemiring
+ }
+
+*-+-isCommutativeSemiring : IsCommutativeSemiring _+_ _*_ 0 1
*-+-isCommutativeSemiring = record
{ +-isCommutativeMonoid = +-0-isCommutativeMonoid
; *-isCommutativeMonoid = *-1-isCommutativeMonoid
@@ -482,7 +669,7 @@ n≤′m+n (suc m) n = ≤′-step (n≤′m+n m n)
; zeroˡ = *-zeroˡ
}
-*-+-commutativeSemiring : CommutativeSemiring _ _
+*-+-commutativeSemiring : CommutativeSemiring 0ℓ 0ℓ
*-+-commutativeSemiring = record
{ isCommutativeSemiring = *-+-isCommutativeSemiring
}
@@ -511,7 +698,7 @@ 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) λ()
+ contradiction (trans (sym $ *-zeroʳ 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)
@@ -526,6 +713,12 @@ i*j≡1⇒j≡1 i j eq = i*j≡1⇒i≡1 j i (trans (*-comm j i) eq)
*-mono-≤ z≤n _ = z≤n
*-mono-≤ (s≤s m≤n) u≤v = +-mono-≤ u≤v (*-mono-≤ m≤n u≤v)
+*-monoˡ-≤ : ∀ n → (_* n) Preserves _≤_ ⟶ _≤_
+*-monoˡ-≤ n m≤o = *-mono-≤ m≤o (≤-refl {n})
+
+*-monoʳ-≤ : ∀ n → (n *_) Preserves _≤_ ⟶ _≤_
+*-monoʳ-≤ n m≤o = *-mono-≤ (≤-refl {n}) m≤o
+
*-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) =
@@ -534,7 +727,7 @@ i*j≡1⇒j≡1 i j eq = i*j≡1⇒i≡1 j i (trans (*-comm j i) eq)
*-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-≤-< (≤-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
@@ -544,6 +737,18 @@ i*j≡1⇒j≡1 i j eq = i*j≡1⇒i≡1 j i (trans (*-comm j i) eq)
------------------------------------------------------------------------
-- Properties of _^_
+^-identityʳ : RightIdentity 1 _^_
+^-identityʳ zero = refl
+^-identityʳ (suc x) = cong suc (^-identityʳ x)
+
+^-zeroˡ : LeftZero 1 _^_
+^-zeroˡ zero = refl
+^-zeroˡ (suc e) = begin
+ 1 ^ suc e ≡⟨⟩
+ 1 * (1 ^ e) ≡⟨ *-identityˡ (1 ^ e) ⟩
+ 1 ^ e ≡⟨ ^-zeroˡ e ⟩
+ 1 ∎
+
^-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
@@ -551,6 +756,29 @@ i*j≡1⇒j≡1 i j eq = i*j≡1⇒i≡1 j i (trans (*-comm j i) eq)
m * ((m ^ n) * (m ^ p)) ≡⟨ sym (*-assoc m _ _) ⟩
(m * (m ^ n)) * (m ^ p) ∎
+^-semigroup-morphism : ∀ {n} → (n ^_) Is +-semigroup -Semigroup⟶ *-semigroup
+^-semigroup-morphism = record
+ { ⟦⟧-cong = cong (_ ^_)
+ ; ∙-homo = ^-distribˡ-+-* _
+ }
+
+^-monoid-morphism : ∀ {n} → (n ^_) Is +-0-monoid -Monoid⟶ *-1-monoid
+^-monoid-morphism = record
+ { sm-homo = ^-semigroup-morphism
+ ; ε-homo = refl
+ }
+
+^-*-assoc : ∀ m n p → (m ^ n) ^ p ≡ m ^ (n * p)
+^-*-assoc m n zero = begin
+ 1 ≡⟨⟩
+ m ^ 0 ≡⟨ cong (m ^_) (sym $ *-zeroʳ n) ⟩
+ m ^ (n * 0) ∎
+^-*-assoc m n (suc p) = begin
+ (m ^ n) * ((m ^ n) ^ p) ≡⟨ cong ((m ^ n) *_) (^-*-assoc m n p) ⟩
+ (m ^ n) * (m ^ (n * p)) ≡⟨ sym (^-distribˡ-+-* m n (n * p)) ⟩
+ m ^ (n + n * p) ≡⟨ cong (m ^_) (sym (+-*-suc n p)) ⟩
+ m ^ (n * (suc 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)
@@ -656,28 +884,43 @@ i^j≡1⇒j≡0∨i≡1 i (suc j) eq = inj₂ (i*j≡1⇒i≡1 i (i ^ j) eq)
⊓-⊔-absorptive : Absorptive _⊓_ _⊔_
⊓-⊔-absorptive = ⊓-abs-⊔ , ⊔-abs-⊓
-⊔-isSemigroup : IsSemigroup _≡_ _⊔_
+⊔-isSemigroup : IsSemigroup _⊔_
⊔-isSemigroup = record
{ isEquivalence = isEquivalence
; assoc = ⊔-assoc
; ∙-cong = cong₂ _⊔_
}
-⊔-0-isCommutativeMonoid : IsCommutativeMonoid _≡_ _⊔_ 0
+⊔-semigroup : Semigroup 0ℓ 0ℓ
+⊔-semigroup = record
+ { isSemigroup = ⊔-isSemigroup
+ }
+
+⊔-0-isCommutativeMonoid : IsCommutativeMonoid _⊔_ 0
⊔-0-isCommutativeMonoid = record
{ isSemigroup = ⊔-isSemigroup
- ; identityˡ = ⊔-identityˡ
+ ; identityˡ = ⊔-identityˡ
; comm = ⊔-comm
}
-⊓-isSemigroup : IsSemigroup _≡_ _⊓_
+⊔-0-commutativeMonoid : CommutativeMonoid 0ℓ 0ℓ
+⊔-0-commutativeMonoid = record
+ { isCommutativeMonoid = ⊔-0-isCommutativeMonoid
+ }
+
+⊓-isSemigroup : IsSemigroup _⊓_
⊓-isSemigroup = record
{ isEquivalence = isEquivalence
; assoc = ⊓-assoc
; ∙-cong = cong₂ _⊓_
}
-⊔-⊓-isSemiringWithoutOne : IsSemiringWithoutOne _≡_ _⊔_ _⊓_ 0
+⊓-semigroup : Semigroup 0ℓ 0ℓ
+⊓-semigroup = record
+ { isSemigroup = ⊔-isSemigroup
+ }
+
+⊔-⊓-isSemiringWithoutOne : IsSemiringWithoutOne _⊔_ _⊓_ 0
⊔-⊓-isSemiringWithoutOne = record
{ +-isCommutativeMonoid = ⊔-0-isCommutativeMonoid
; *-isSemigroup = ⊓-isSemigroup
@@ -686,19 +929,19 @@ i^j≡1⇒j≡0∨i≡1 i (suc j) eq = inj₂ (i*j≡1⇒i≡1 i (i ^ j) eq)
}
⊔-⊓-isCommutativeSemiringWithoutOne
- : IsCommutativeSemiringWithoutOne _≡_ _⊔_ _⊓_ 0
+ : IsCommutativeSemiringWithoutOne _⊔_ _⊓_ 0
⊔-⊓-isCommutativeSemiringWithoutOne = record
{ isSemiringWithoutOne = ⊔-⊓-isSemiringWithoutOne
; *-comm = ⊓-comm
}
-⊔-⊓-commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne _ _
+⊔-⊓-commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne 0ℓ 0ℓ
⊔-⊓-commutativeSemiringWithoutOne = record
{ isCommutativeSemiringWithoutOne =
⊔-⊓-isCommutativeSemiringWithoutOne
}
-⊓-⊔-isLattice : IsLattice _≡_ _⊓_ _⊔_
+⊓-⊔-isLattice : IsLattice _⊓_ _⊔_
⊓-⊔-isLattice = record
{ isEquivalence = isEquivalence
; ∨-comm = ⊓-comm
@@ -710,13 +953,18 @@ i^j≡1⇒j≡0∨i≡1 i (suc j) eq = inj₂ (i*j≡1⇒i≡1 i (i ^ j) eq)
; absorptive = ⊓-⊔-absorptive
}
-⊓-⊔-isDistributiveLattice : IsDistributiveLattice _≡_ _⊓_ _⊔_
+⊓-⊔-lattice : Lattice 0ℓ 0ℓ
+⊓-⊔-lattice = record
+ { isLattice = ⊓-⊔-isLattice
+ }
+
+⊓-⊔-isDistributiveLattice : IsDistributiveLattice _⊓_ _⊔_
⊓-⊔-isDistributiveLattice = record
{ isLattice = ⊓-⊔-isLattice
; ∨-∧-distribʳ = ⊓-distribʳ-⊔
}
-⊓-⊔-distributiveLattice : DistributiveLattice _ _
+⊓-⊔-distributiveLattice : DistributiveLattice 0ℓ 0ℓ
⊓-⊔-distributiveLattice = record
{ isDistributiveLattice = ⊓-⊔-isDistributiveLattice
}
@@ -743,11 +991,43 @@ 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)
+m≤n⇒m⊓n≡m : ∀ {m n} → m ≤ n → m ⊓ n ≡ m
+m≤n⇒m⊓n≡m z≤n = refl
+m≤n⇒m⊓n≡m (s≤s m≤n) = cong suc (m≤n⇒m⊓n≡m m≤n)
+
+m≤n⇒n⊓m≡m : ∀ {m n} → m ≤ n → n ⊓ m ≡ m
+m≤n⇒n⊓m≡m {m} m≤n = trans (⊓-comm _ m) (m≤n⇒m⊓n≡m m≤n)
+
+m⊓n≡m⇒m≤n : ∀ {m n} → m ⊓ n ≡ m → m ≤ n
+m⊓n≡m⇒m≤n m⊓n≡m = subst (_≤ _) m⊓n≡m (m⊓n≤n _ _)
+
+m⊓n≡n⇒n≤m : ∀ {m n} → m ⊓ n ≡ n → n ≤ m
+m⊓n≡n⇒n≤m m⊓n≡n = subst (_≤ _) m⊓n≡n (m⊓n≤m _ _)
+
+m≤n⇒n⊔m≡n : ∀ {m n} → m ≤ n → n ⊔ m ≡ n
+m≤n⇒n⊔m≡n z≤n = ⊔-identityʳ _
+m≤n⇒n⊔m≡n (s≤s m≤n) = cong suc (m≤n⇒n⊔m≡n m≤n)
+
+m≤n⇒m⊔n≡n : ∀ {m n} → m ≤ n → m ⊔ n ≡ n
+m≤n⇒m⊔n≡n {m} m≤n = trans (⊔-comm m _) (m≤n⇒n⊔m≡n m≤n)
+
+n⊔m≡m⇒n≤m : ∀ {m n} → n ⊔ m ≡ m → n ≤ m
+n⊔m≡m⇒n≤m n⊔m≡m = subst (_ ≤_) n⊔m≡m (m≤m⊔n _ _)
+
+n⊔m≡n⇒m≤n : ∀ {m n} → n ⊔ m ≡ n → m ≤ n
+n⊔m≡n⇒m≤n n⊔m≡n = subst (_ ≤_) n⊔m≡n (n≤m⊔n _ _)
+
⊔-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)
+⊔-monoˡ-≤ : ∀ n → (_⊔ n) Preserves _≤_ ⟶ _≤_
+⊔-monoˡ-≤ n m≤o = ⊔-mono-≤ m≤o (≤-refl {n})
+
+⊔-monoʳ-≤ : ∀ n → (n ⊔_) Preserves _≤_ ⟶ _≤_
+⊔-monoʳ-≤ n m≤o = ⊔-mono-≤ (≤-refl {n}) m≤o
+
⊔-mono-< : _⊔_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
⊔-mono-< = ⊔-mono-≤
@@ -756,6 +1036,12 @@ m⊓n≤m⊔n (suc m) (suc n) = s≤s (m⊓n≤m⊔n m n)
... | 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
+⊓-monoˡ-≤ : ∀ n → (_⊓ n) Preserves _≤_ ⟶ _≤_
+⊓-monoˡ-≤ n m≤o = ⊓-mono-≤ m≤o (≤-refl {n})
+
+⊓-monoʳ-≤ : ∀ n → (n ⊓_) Preserves _≤_ ⟶ _≤_
+⊓-monoʳ-≤ n m≤o = ⊓-mono-≤ (≤-refl {n}) m≤o
+
⊓-mono-< : _⊓_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_
⊓-mono-< = ⊓-mono-≤
@@ -790,6 +1076,23 @@ m⊓n≤m+n m n with ⊓-sel m n
+-distrib-⊓ : _+_ DistributesOver _⊓_
+-distrib-⊓ = +-distribˡ-⊓ , +-distribʳ-⊓
+-- Other properties
+⊓-triangulate : ∀ x y z → x ⊓ y ⊓ z ≡ (x ⊓ y) ⊓ (y ⊓ z)
+⊓-triangulate x y z = begin
+ x ⊓ y ⊓ z ≡⟨ cong (λ v → x ⊓ v ⊓ z) (sym (⊓-idem y)) ⟩
+ x ⊓ (y ⊓ y) ⊓ z ≡⟨ ⊓-assoc x _ _ ⟩
+ x ⊓ ((y ⊓ y) ⊓ z) ≡⟨ cong (x ⊓_) (⊓-assoc y _ _) ⟩
+ x ⊓ (y ⊓ (y ⊓ z)) ≡⟨ sym (⊓-assoc x _ _) ⟩
+ (x ⊓ y) ⊓ (y ⊓ z) ∎
+
+⊔-triangulate : ∀ x y z → x ⊔ y ⊔ z ≡ (x ⊔ y) ⊔ (y ⊔ z)
+⊔-triangulate x y z = begin
+ x ⊔ y ⊔ z ≡⟨ cong (λ v → x ⊔ v ⊔ z) (sym (⊔-idem y)) ⟩
+ x ⊔ (y ⊔ y) ⊔ z ≡⟨ ⊔-assoc x _ _ ⟩
+ x ⊔ ((y ⊔ y) ⊔ z) ≡⟨ cong (x ⊔_) (⊔-assoc y _ _) ⟩
+ x ⊔ (y ⊔ (y ⊔ z)) ≡⟨ sym (⊔-assoc x _ _) ⟩
+ (x ⊔ y) ⊔ (y ⊔ z) ∎
+
------------------------------------------------------------------------
-- Properties of _∸_
@@ -801,6 +1104,38 @@ n∸n≡0 : ∀ n → n ∸ n ≡ 0
n∸n≡0 zero = refl
n∸n≡0 (suc n) = n∸n≡0 n
+-- Ordering properties of _∸_
+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)
+
+m≮m∸n : ∀ m n → m ≮ m ∸ n
+m≮m∸n zero (suc n) ()
+m≮m∸n m zero = n≮n m
+m≮m∸n (suc m) (suc n) = m≮m∸n m n ∘ ≤-trans (n≤1+n (suc m))
+
+∸-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₂
+
+∸-monoˡ-≤ : ∀ {m n} o → m ≤ n → m ∸ o ≤ n ∸ o
+∸-monoˡ-≤ o m≤n = ∸-mono {u = o} m≤n ≤-refl
+
+∸-monoʳ-≤ : ∀ {m n} o → m ≤ n → o ∸ m ≥ o ∸ n
+∸-monoʳ-≤ _ m≤n = ∸-mono ≤-refl m≤n
+
+m∸n≡0⇒m≤n : ∀ {m n} → m ∸ n ≡ 0 → m ≤ n
+m∸n≡0⇒m≤n {zero} {_} _ = z≤n
+m∸n≡0⇒m≤n {suc m} {zero} ()
+m∸n≡0⇒m≤n {suc m} {suc n} eq = s≤s (m∸n≡0⇒m≤n eq)
+
+m≤n⇒m∸n≡0 : ∀ {m n} → m ≤ n → m ∸ n ≡ 0
+m≤n⇒m∸n≡0 {n = n} z≤n = 0∸n≡0 n
+m≤n⇒m∸n≡0 {_} (s≤s m≤n) = m≤n⇒m∸n≡0 m≤n
+
+-- Properties of _∸_ and _+_
+-∸-comm : ∀ {m} n {o} → o ≤ m → (m + n) ∸ o ≡ (m ∸ o) + n
+-∸-comm {zero} _ {suc o} ()
+-∸-comm {zero} _ {zero} _ = refl
@@ -822,11 +1157,6 @@ n∸n≡0 (suc n) = n∸n≡0 n
(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
@@ -846,6 +1176,41 @@ m+n∸m≡n {m} {n} m≤n = begin
(n + m) ∸ m ≡⟨ m+n∸n≡m n m ⟩
n ∎
+m∸n+n≡m : ∀ {m n} → n ≤ m → (m ∸ n) + n ≡ m
+m∸n+n≡m {m} {n} n≤m = begin
+ (m ∸ n) + n ≡⟨ sym (+-∸-comm n n≤m) ⟩
+ (m + n) ∸ n ≡⟨ m+n∸n≡m m n ⟩
+ m ∎
+
+m∸[m∸n]≡n : ∀ {m n} → n ≤ m → m ∸ (m ∸ n) ≡ n
+m∸[m∸n]≡n {m} {_} z≤n = n∸n≡0 m
+m∸[m∸n]≡n {suc m} {suc n} (s≤s n≤m) = begin
+ suc m ∸ (m ∸ n) ≡⟨ +-∸-assoc 1 (n∸m≤n n m) ⟩
+ suc (m ∸ (m ∸ n)) ≡⟨ cong suc (m∸[m∸n]≡n n≤m) ⟩
+ suc n ∎
+
+[i+j]∸[i+k]≡j∸k : ∀ i j k → (i + j) ∸ (i + k) ≡ j ∸ k
+[i+j]∸[i+k]≡j∸k zero j k = refl
+[i+j]∸[i+k]≡j∸k (suc i) j k = [i+j]∸[i+k]≡j∸k i j k
+
+-- Properties of _∸_ and _*_
+*-distribʳ-∸ : _*_ DistributesOverʳ _∸_
+*-distribʳ-∸ i zero zero = refl
+*-distribʳ-∸ zero zero (suc k) = sym (0∸n≡0 (k * zero))
+*-distribʳ-∸ (suc i) zero (suc k) = refl
+*-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) ∎
+
+*-distribˡ-∸ : _*_ DistributesOverˡ _∸_
+*-distribˡ-∸ = comm+distrʳ⇒distrˡ (cong₂ _∸_) *-comm *-distribʳ-∸
+
+*-distrib-∸ : _*_ DistributesOver _∸_
+*-distrib-∸ = *-distribˡ-∸ , *-distribʳ-∸
+
+-- Properties of _∸_ and _⊓_ and _⊔_
m⊓n+n∸m≡n : ∀ m n → (m ⊓ n) + (n ∸ m) ≡ n
m⊓n+n∸m≡n zero n = refl
m⊓n+n∸m≡n (suc m) zero = refl
@@ -857,40 +1222,14 @@ m⊓n+n∸m≡n (suc m) (suc n) = cong suc $ m⊓n+n∸m≡n m n
[m∸n]⊓[n∸m]≡0 (suc m) zero = refl
[m∸n]⊓[n∸m]≡0 (suc m) (suc n) = [m∸n]⊓[n∸m]≡0 m n
-[i+j]∸[i+k]≡j∸k : ∀ i j k → (i + j) ∸ (i + k) ≡ j ∸ k
-[i+j]∸[i+k]≡j∸k zero j k = refl
-[i+j]∸[i+k]≡j∸k (suc i) j k = [i+j]∸[i+k]≡j∸k i j k
-
--- 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 (_+ (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 ≡⟨⟩
- suc (i + j) ∎
-i∸k∸j+j∸k≡i+j∸k (suc i) zero (suc k) = begin
- 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 (_∸ k) (sym (+-suc i j)) ⟩
- i + suc j ∸ k ∎
-
-*-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 = 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) ∎
+∸-distribˡ-⊓-⊔ : ∀ x y z → x ∸ (y ⊓ z) ≡ (x ∸ y) ⊔ (x ∸ z)
+∸-distribˡ-⊓-⊔ x zero zero = sym (⊔-idem x)
+∸-distribˡ-⊓-⊔ zero zero (suc z) = refl
+∸-distribˡ-⊓-⊔ zero (suc y) zero = refl
+∸-distribˡ-⊓-⊔ zero (suc y) (suc z) = refl
+∸-distribˡ-⊓-⊔ (suc x) (suc y) zero = sym (m≤n⇒m⊔n≡n (≤-step (n∸m≤n y x)))
+∸-distribˡ-⊓-⊔ (suc x) zero (suc z) = sym (m≤n⇒n⊔m≡n (≤-step (n∸m≤n z x)))
+∸-distribˡ-⊓-⊔ (suc x) (suc y) (suc z) = ∸-distribˡ-⊓-⊔ x y z
∸-distribʳ-⊓ : _∸_ DistributesOverʳ _⊓_
∸-distribʳ-⊓ zero y z = refl
@@ -898,24 +1237,102 @@ i∸k∸j+j∸k≡i+j∸k (suc i) (suc j) (suc k) = begin
∸-distribʳ-⊓ (suc x) (suc y) zero = sym (⊓-zeroʳ (y ∸ x))
∸-distribʳ-⊓ (suc x) (suc y) (suc z) = ∸-distribʳ-⊓ x y z
+∸-distribˡ-⊔-⊓ : ∀ x y z → x ∸ (y ⊔ z) ≡ (x ∸ y) ⊓ (x ∸ z)
+∸-distribˡ-⊔-⊓ x zero zero = sym (⊓-idem x)
+∸-distribˡ-⊔-⊓ zero zero z = 0∸n≡0 z
+∸-distribˡ-⊔-⊓ zero (suc y) z = 0∸n≡0 (suc y ⊔ z)
+∸-distribˡ-⊔-⊓ (suc x) (suc y) zero = sym (m≤n⇒m⊓n≡m (≤-step (n∸m≤n y x)))
+∸-distribˡ-⊔-⊓ (suc x) zero (suc z) = sym (m≤n⇒n⊓m≡m (≤-step (n∸m≤n z 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 (_∸ 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 ∎
-
-∸-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 ∣_-_∣
+
+n≡m⇒∣n-m∣≡0 : ∀ {n m} → n ≡ m → ∣ n - m ∣ ≡ 0
+n≡m⇒∣n-m∣≡0 {zero} refl = refl
+n≡m⇒∣n-m∣≡0 {suc n} refl = n≡m⇒∣n-m∣≡0 {n} refl
+
+m≤n⇒∣n-m∣≡n∸m : ∀ {n m} → m ≤ n → ∣ n - m ∣ ≡ n ∸ m
+m≤n⇒∣n-m∣≡n∸m {zero} z≤n = refl
+m≤n⇒∣n-m∣≡n∸m {suc n} z≤n = refl
+m≤n⇒∣n-m∣≡n∸m (s≤s m≤n) = m≤n⇒∣n-m∣≡n∸m m≤n
+
+∣n-m∣≡0⇒n≡m : ∀ {n m} → ∣ n - m ∣ ≡ 0 → n ≡ m
+∣n-m∣≡0⇒n≡m {zero} {zero} eq = refl
+∣n-m∣≡0⇒n≡m {zero} {suc m} ()
+∣n-m∣≡0⇒n≡m {suc n} {zero} ()
+∣n-m∣≡0⇒n≡m {suc n} {suc m} eq = cong suc (∣n-m∣≡0⇒n≡m eq)
+
+∣n-m∣≡n∸m⇒m≤n : ∀ {n m} → ∣ n - m ∣ ≡ n ∸ m → m ≤ n
+∣n-m∣≡n∸m⇒m≤n {zero} {zero} eq = z≤n
+∣n-m∣≡n∸m⇒m≤n {zero} {suc m} ()
+∣n-m∣≡n∸m⇒m≤n {suc n} {zero} eq = z≤n
+∣n-m∣≡n∸m⇒m≤n {suc n} {suc m} eq = s≤s (∣n-m∣≡n∸m⇒m≤n eq)
+
+∣n-n∣≡0 : ∀ n → ∣ n - n ∣ ≡ 0
+∣n-n∣≡0 n = n≡m⇒∣n-m∣≡0 {n} refl
+
+∣n-n+m∣≡m : ∀ n m → ∣ n - n + m ∣ ≡ m
+∣n-n+m∣≡m zero m = refl
+∣n-n+m∣≡m (suc n) m = ∣n-n+m∣≡m n m
+
+∣n+m-n+o∣≡∣m-o| : ∀ n m o → ∣ n + m - n + o ∣ ≡ ∣ m - o ∣
+∣n+m-n+o∣≡∣m-o| zero m o = refl
+∣n+m-n+o∣≡∣m-o| (suc n) m o = ∣n+m-n+o∣≡∣m-o| n m o
+
+n∸m≤∣n-m∣ : ∀ n m → n ∸ m ≤ ∣ n - m ∣
+n∸m≤∣n-m∣ n m with ≤-total m n
+... | inj₁ m≤n = subst (n ∸ m ≤_) (sym (m≤n⇒∣n-m∣≡n∸m m≤n)) ≤-refl
+... | inj₂ n≤m = subst (_≤ ∣ n - m ∣) (sym (m≤n⇒m∸n≡0 n≤m)) z≤n
+
+∣n-m∣≤n⊔m : ∀ n m → ∣ n - m ∣ ≤ n ⊔ m
+∣n-m∣≤n⊔m zero m = ≤-refl
+∣n-m∣≤n⊔m (suc n) zero = ≤-refl
+∣n-m∣≤n⊔m (suc n) (suc m) = ≤-step (∣n-m∣≤n⊔m n m)
+
+∣-∣-comm : Commutative ∣_-_∣
+∣-∣-comm zero zero = refl
+∣-∣-comm zero (suc m) = refl
+∣-∣-comm (suc n) zero = refl
+∣-∣-comm (suc n) (suc m) = ∣-∣-comm n m
+
+∣n-m∣≡[n∸m]∨[m∸n] : ∀ m n → (∣ n - m ∣ ≡ n ∸ m) ⊎ (∣ n - m ∣ ≡ m ∸ n)
+∣n-m∣≡[n∸m]∨[m∸n] m n with ≤-total m n
+... | inj₁ m≤n = inj₁ $ m≤n⇒∣n-m∣≡n∸m m≤n
+... | inj₂ n≤m = inj₂ $ begin
+ ∣ n - m ∣ ≡⟨ ∣-∣-comm n m ⟩
+ ∣ m - n ∣ ≡⟨ m≤n⇒∣n-m∣≡n∸m n≤m ⟩
+ m ∸ n ∎
+
+private
+
+ *-distribˡ-∣-∣-aux : ∀ a m n → m ≤ n → a * ∣ n - m ∣ ≡ ∣ a * n - a * m ∣
+ *-distribˡ-∣-∣-aux a m n m≤n = begin
+ a * ∣ n - m ∣ ≡⟨ cong (a *_) (m≤n⇒∣n-m∣≡n∸m m≤n) ⟩
+ a * (n ∸ m) ≡⟨ *-distribˡ-∸ a n m ⟩
+ a * n ∸ a * m ≡⟨ sym $′ m≤n⇒∣n-m∣≡n∸m (*-monoʳ-≤ a m≤n) ⟩
+ ∣ a * n - a * m ∣ ∎
+
+*-distribˡ-∣-∣ : _*_ DistributesOverˡ ∣_-_∣
+*-distribˡ-∣-∣ a m n with ≤-total m n
+... | inj₁ m≤n = begin
+ a * ∣ m - n ∣ ≡⟨ cong (a *_) (∣-∣-comm m n) ⟩
+ a * ∣ n - m ∣ ≡⟨ *-distribˡ-∣-∣-aux a m n m≤n ⟩
+ ∣ a * n - a * m ∣ ≡⟨ ∣-∣-comm (a * n) (a * m) ⟩
+ ∣ a * m - a * n ∣ ∎
+... | inj₂ n≤m = *-distribˡ-∣-∣-aux a n m n≤m
+
+*-distribʳ-∣-∣ : _*_ DistributesOverʳ ∣_-_∣
+*-distribʳ-∣-∣ = comm+distrˡ⇒distrʳ (cong₂ ∣_-_∣) *-comm *-distribˡ-∣-∣
+
+*-distrib-∣-∣ : _*_ DistributesOver ∣_-_∣
+*-distrib-∣-∣ = *-distribˡ-∣-∣ , *-distribʳ-∣-∣
------------------------------------------------------------------------
-- Properties of ⌊_/2⌋
@@ -938,17 +1355,22 @@ im≡jm+n⇒[i∸j]m≡n i j m n eq = begin
⌊n/2⌋≤′n (suc n) = ≤′-step (⌈n/2⌉≤′n n)
------------------------------------------------------------------------
--- Modules for reasoning about natural number relations
+-- Other properties
+
+-- If there is an injection from a type to ℕ, then the type has
+-- decidable equality.
--- A module for automatically solving propositional equivalences
-module SemiringSolver =
- Solver (ACR.fromCommutativeSemiring *-+-commutativeSemiring) _≟_
+eq? : ∀ {a} {A : Set a} → A ↣ ℕ → Decidable {A = A} _≡_
+eq? inj = via-injection inj _≟_
+
+------------------------------------------------------------------------
+-- Modules for reasoning about natural number relations
-- A module for reasoning about the _≤_ relation
module ≤-Reasoning where
open import Relation.Binary.PartialOrderReasoning
(DecTotalOrder.poset ≤-decTotalOrder) public
- renaming (_≈⟨_⟩_ to _≡⟨_⟩_)
+ hiding (_≈⟨_⟩_)
infixr 2 _<⟨_⟩_
@@ -961,23 +1383,140 @@ module ≤-Reasoning where
-- Please use the new names as continuing support for the old names is
-- not guaranteed.
+-- Version 0.14
+
_*-mono_ = *-mono-≤
+{-# WARNING_ON_USAGE _*-mono_
+"Warning: _*-mono_ was deprecated in v0.14.
+Please use *-mono-≤ instead."
+#-}
_+-mono_ = +-mono-≤
-
+{-# WARNING_ON_USAGE _+-mono_
+"Warning: _+-mono_ was deprecated in v0.14.
+Please use +-mono-≤ instead."
+#-}
+-right-identity = +-identityʳ
+{-# WARNING_ON_USAGE +-right-identity
+"Warning: +-right-identity was deprecated in v0.14.
+Please use +-identityʳ instead."
+#-}
*-right-zero = *-zeroʳ
+{-# WARNING_ON_USAGE *-right-zero
+"Warning: *-right-zero was deprecated in v0.14.
+Please use *-zeroʳ instead."
+#-}
distribʳ-*-+ = *-distribʳ-+
+{-# WARNING_ON_USAGE distribʳ-*-+
+"Warning: distribʳ-*-+ was deprecated in v0.14.
+Please use *-distribʳ-+ instead."
+#-}
*-distrib-∸ʳ = *-distribʳ-∸
+{-# WARNING_ON_USAGE *-distrib-∸ʳ
+"Warning: *-distrib-∸ʳ was deprecated in v0.14.
+Please use *-distribʳ-∸ instead."
+#-}
cancel-+-left = +-cancelˡ-≡
+{-# WARNING_ON_USAGE cancel-+-left
+"Warning: cancel-+-left was deprecated in v0.14.
+Please use +-cancelˡ-≡ instead."
+#-}
cancel-+-left-≤ = +-cancelˡ-≤
+{-# WARNING_ON_USAGE cancel-+-left-≤
+"Warning: cancel-+-left-≤ was deprecated in v0.14.
+Please use +-cancelˡ-≤ instead."
+#-}
cancel-*-right = *-cancelʳ-≡
+{-# WARNING_ON_USAGE cancel-*-right
+"Warning: cancel-*-right was deprecated in v0.14.
+Please use *-cancelʳ-≡ instead."
+#-}
cancel-*-right-≤ = *-cancelʳ-≤
-
+{-# WARNING_ON_USAGE cancel-*-right-≤
+"Warning: cancel-*-right-≤ was deprecated in v0.14.
+Please use *-cancelʳ-≤ instead."
+#-}
strictTotalOrder = <-strictTotalOrder
+{-# WARNING_ON_USAGE strictTotalOrder
+"Warning: strictTotalOrder was deprecated in v0.14.
+Please use <-strictTotalOrder instead."
+#-}
isCommutativeSemiring = *-+-isCommutativeSemiring
+{-# WARNING_ON_USAGE isCommutativeSemiring
+"Warning: isCommutativeSemiring was deprecated in v0.14.
+Please use *-+-isCommutativeSemiring instead."
+#-}
commutativeSemiring = *-+-commutativeSemiring
+{-# WARNING_ON_USAGE commutativeSemiring
+"Warning: commutativeSemiring was deprecated in v0.14.
+Please use *-+-commutativeSemiring instead."
+#-}
isDistributiveLattice = ⊓-⊔-isDistributiveLattice
+{-# WARNING_ON_USAGE isDistributiveLattice
+"Warning: isDistributiveLattice was deprecated in v0.14.
+Please use ⊓-⊔-isDistributiveLattice instead."
+#-}
distributiveLattice = ⊓-⊔-distributiveLattice
+{-# WARNING_ON_USAGE distributiveLattice
+"Warning: distributiveLattice was deprecated in v0.14.
+Please use ⊓-⊔-distributiveLattice instead."
+#-}
⊔-⊓-0-isSemiringWithoutOne = ⊔-⊓-isSemiringWithoutOne
+{-# WARNING_ON_USAGE ⊔-⊓-0-isSemiringWithoutOne
+"Warning: ⊔-⊓-0-isSemiringWithoutOne was deprecated in v0.14.
+Please use ⊔-⊓-isSemiringWithoutOne instead."
+#-}
⊔-⊓-0-isCommutativeSemiringWithoutOne = ⊔-⊓-isCommutativeSemiringWithoutOne
+{-# WARNING_ON_USAGE ⊔-⊓-0-isCommutativeSemiringWithoutOne
+"Warning: ⊔-⊓-0-isCommutativeSemiringWithoutOne was deprecated in v0.14.
+Please use ⊔-⊓-isCommutativeSemiringWithoutOne instead."
+#-}
⊔-⊓-0-commutativeSemiringWithoutOne = ⊔-⊓-commutativeSemiringWithoutOne
+{-# WARNING_ON_USAGE ⊔-⊓-0-commutativeSemiringWithoutOne
+"Warning: ⊔-⊓-0-commutativeSemiringWithoutOne was deprecated in v0.14.
+Please use ⊔-⊓-commutativeSemiringWithoutOne instead."
+#-}
+
+-- Version 0.15
+
+¬i+1+j≤i = i+1+j≰i
+{-# WARNING_ON_USAGE ¬i+1+j≤i
+"Warning: ¬i+1+j≤i was deprecated in v0.15.
+Please use i+1+j≰i instead."
+#-}
+≤-steps = ≤-stepsˡ
+{-# WARNING_ON_USAGE ≤-steps
+"Warning: ≤-steps was deprecated in v0.15.
+Please use ≤-stepsˡ instead."
+#-}
+
+-- Version 0.17
+
+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 = cong (_+ (j ∸ k)) (0∸n≡0 (k ∸ j))
+i∸k∸j+j∸k≡i+j∸k (suc i) j zero = cong (λ x → suc i ∸ x + j) (0∸n≡0 j)
+i∸k∸j+j∸k≡i+j∸k (suc i) zero (suc k) = begin
+ 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 (_∸ k) (sym (+-suc i j)) ⟩
+ i + suc j ∸ k ∎
+{-# WARNING_ON_USAGE i∸k∸j+j∸k≡i+j∸k
+"Warning: i∸k∸j+j∸k≡i+j∸k was deprecated in v0.17."
+#-}
+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 (_∸ 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 ∎
+{-# WARNING_ON_USAGE im≡jm+n⇒[i∸j]m≡n
+"Warning: im≡jm+n⇒[i∸j]m≡n was deprecated in v0.17."
+#-}
+≤+≢⇒< = ≤∧≢⇒<
+{-# WARNING_ON_USAGE ≤+≢⇒<
+"Warning: ≤+≢⇒< was deprecated in v0.17.
+Please use ≤∧≢⇒< instead."
+#-}
diff --git a/src/Data/Nat/Solver.agda b/src/Data/Nat/Solver.agda
new file mode 100644
index 0000000..21a8e33
--- /dev/null
+++ b/src/Data/Nat/Solver.agda
@@ -0,0 +1,21 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Automatic solvers for equations over naturals
+------------------------------------------------------------------------
+
+-- See README.Nat for examples of how to use this solver
+
+module Data.Nat.Solver where
+
+import Algebra.Solver.Ring.Simple as Solver
+import Algebra.Solver.Ring.AlmostCommutativeRing as ACR
+open import Data.Nat using (_≟_)
+open import Data.Nat.Properties
+
+------------------------------------------------------------------------
+-- A module for automatically solving propositional equivalences
+-- containing _+_ and _*_
+
+module +-*-Solver =
+ Solver (ACR.fromCommutativeSemiring *-+-commutativeSemiring) _≟_
diff --git a/src/Data/Nat/Unsafe.agda b/src/Data/Nat/Unsafe.agda
new file mode 100644
index 0000000..34839cf
--- /dev/null
+++ b/src/Data/Nat/Unsafe.agda
@@ -0,0 +1,13 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Unsafe natural number types and operations
+------------------------------------------------------------------------
+
+module Data.Nat.Unsafe where
+
+open import Data.Nat.Base
+import Relation.Binary.PropositionalEquality.TrustMe as TrustMe
+
+erase : ∀ {m n} → m ≤″ n → m ≤″ n
+erase (less-than-or-equal eq) = less-than-or-equal (TrustMe.erase eq)
diff --git a/src/Data/Plus.agda b/src/Data/Plus.agda
index 8d85539..663b8a7 100644
--- a/src/Data/Plus.agda
+++ b/src/Data/Plus.agda
@@ -2,83 +2,11 @@
-- The Agda standard library
--
-- Transitive closures
-------------------------------------------------------------------------
-
-module Data.Plus where
-
-open import Function
-open import Function.Equivalence as Equiv using (_⇔_)
-open import Level
-open import Relation.Binary
-
-------------------------------------------------------------------------
--- Transitive closure
-
-infix 4 Plus
-
-syntax Plus R x y = x [ R ]⁺ y
-
-data Plus {a ℓ} {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where
- [_] : ∀ {x y} (x∼y : x ∼ y) → x [ _∼_ ]⁺ y
- _∼⁺⟨_⟩_ : ∀ x {y z} (x∼⁺y : x [ _∼_ ]⁺ y) (y∼⁺z : y [ _∼_ ]⁺ z) →
- x [ _∼_ ]⁺ z
-
--- "Equational" reasoning notation. Example:
--
--- lemma =
--- x ∼⁺⟨ [ lemma₁ ] ⟩
--- y ∼⁺⟨ lemma₂ ⟩∎
--- z ∎
-
-finally : ∀ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} x y →
- x [ _∼_ ]⁺ y → x [ _∼_ ]⁺ y
-finally _ _ = id
-
-syntax finally x y x∼⁺y = x ∼⁺⟨ x∼⁺y ⟩∎ y ∎
-
-infixr 2 _∼⁺⟨_⟩_
-infix 3 finally
-
--- Map.
-
-map : ∀ {a a′ ℓ ℓ′} {A : Set a} {A′ : Set a′}
- {_R_ : Rel A ℓ} {_R′_ : Rel A′ ℓ′} {f : A → A′} →
- _R_ =[ f ]⇒ _R′_ → Plus _R_ =[ f ]⇒ Plus _R′_
-map R⇒R′ [ xRy ] = [ R⇒R′ xRy ]
-map R⇒R′ (x ∼⁺⟨ xR⁺z ⟩ zR⁺y) =
- _ ∼⁺⟨ map R⇒R′ xR⁺z ⟩ map R⇒R′ zR⁺y
-
+-- This module is DEPRECATED. Please use the
+-- Relation.Binary.Construct.Closure.Transitive module directly.
------------------------------------------------------------------------
--- Alternative definition of transitive closure
-
--- A generalisation of Data.List.Nonempty.List⁺.
-
-infixr 5 _∷_ _++_
-infix 4 Plus′
-syntax Plus′ R x y = x ⟨ R ⟩⁺ y
-
-data Plus′ {a ℓ} {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where
- [_] : ∀ {x y} (x∼y : x ∼ y) → x ⟨ _∼_ ⟩⁺ y
- _∷_ : ∀ {x y z} (x∼y : x ∼ y) (y∼⁺z : y ⟨ _∼_ ⟩⁺ z) → x ⟨ _∼_ ⟩⁺ z
-
--- Transitivity.
-
-_++_ : ∀ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} {x y z} →
- x ⟨ _∼_ ⟩⁺ y → y ⟨ _∼_ ⟩⁺ z → x ⟨ _∼_ ⟩⁺ z
-[ x∼y ] ++ y∼⁺z = x∼y ∷ y∼⁺z
-(x∼y ∷ y∼⁺z) ++ z∼⁺u = x∼y ∷ (y∼⁺z ++ z∼⁺u)
-
--- Plus and Plus′ are equivalent.
-
-equivalent : ∀ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} {x y} →
- Plus _∼_ x y ⇔ Plus′ _∼_ x y
-equivalent {_∼_ = _∼_} = Equiv.equivalence complete sound
- where
- complete : Plus _∼_ ⇒ Plus′ _∼_
- complete [ x∼y ] = [ x∼y ]
- complete (x ∼⁺⟨ x∼⁺y ⟩ y∼⁺z) = complete x∼⁺y ++ complete y∼⁺z
+module Data.Plus where
- sound : Plus′ _∼_ ⇒ Plus _∼_
- sound [ x∼y ] = [ x∼y ]
- sound (x∼y ∷ y∼⁺z) = _ ∼⁺⟨ [ x∼y ] ⟩ sound y∼⁺z
+open import Relation.Binary.Construct.Closure.Transitive public
diff --git a/src/Data/Product.agda b/src/Data/Product.agda
index bfc8939..953bfa9 100644
--- a/src/Data/Product.agda
+++ b/src/Data/Product.agda
@@ -9,21 +9,18 @@ module Data.Product where
open import Function
open import Level
open import Relation.Nullary
+open import Agda.Builtin.Equality
-infixr 4 _,_ _,′_
-infix 4 ,_
+infixr 4 _,′_
+infix 4 -,_
infixr 2 _×_ _-×-_ _-,-_
------------------------------------------------------------------------
-- Definition
-record Σ {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
- constructor _,_
- field
- proj₁ : A
- proj₂ : B proj₁
+open import Agda.Builtin.Sigma hiding (module Σ) public renaming (fst to proj₁; snd to proj₂)
-open Σ public
+module Σ = Agda.Builtin.Sigma.Σ renaming (fst to proj₁; snd to proj₂)
-- The syntax declaration below is attached to Σ-syntax, to make it
-- easy to import Σ without the special syntax.
@@ -75,8 +72,8 @@ _,′_ = _,_
-- Sometimes the first component can be inferred.
-,_ : ∀ {a b} {A : Set a} {B : A → Set b} {x} → B x → ∃ B
-, y = _ , y
+-,_ : ∀ {a b} {A : Set a} {B : A → Set b} {x} → B x → ∃ B
+-, y = _ , y
<_,_> : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ {x} → B x → Set c}
(f : (x : A) → B x) → ((x : A) → C (f x)) →
@@ -84,11 +81,19 @@ _,′_ = _,_
< f , g > x = (f x , g x)
map : ∀ {a b p q}
- {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q} →
+ {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q} →
(f : A → B) → (∀ {x} → P x → Q (f x)) →
Σ A P → Σ B Q
map f g (x , y) = (f x , g y)
+map₁ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
+ (A → B) → A × C → B × C
+map₁ f = map f id
+
+map₂ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : A → Set c} →
+ (∀ {x} → B x → C x) → Σ A B → Σ A C
+map₂ f = map id f
+
zip : ∀ {a b c p q r}
{A : Set a} {B : Set b} {C : Set c}
{P : A → Set p} {Q : B → Set q} {R : C → Set r} →
@@ -113,6 +118,10 @@ curry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} →
((x : A) → (y : B x) → C (x , y))
curry f x y = f (x , y)
+curry′ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
+ (A × B → C) → (A → B → C)
+curry′ = curry
+
uncurry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} →
((x : A) → (y : B x) → C (x , y)) →
((p : Σ A B) → C p)
diff --git a/src/Data/Product/Categorical/Examples.agda b/src/Data/Product/Categorical/Examples.agda
new file mode 100644
index 0000000..26eba74
--- /dev/null
+++ b/src/Data/Product/Categorical/Examples.agda
@@ -0,0 +1,60 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Universe-sensitive functor and monad instances for the Product type.
+------------------------------------------------------------------------
+
+open import Algebra
+
+module Data.Product.Categorical.Examples
+ {a e b} {A : Monoid a e} {B : Set b} where
+
+open import Level using (Lift; lift; _⊔_)
+open import Category.Functor using (RawFunctor)
+open import Category.Monad using (RawMonad)
+open import Data.Product
+open import Data.Product.Relation.Pointwise.NonDependent
+open import Function
+import Function.Identity.Categorical as Id
+open import Relation.Binary using (Rel)
+open import Relation.Binary.PropositionalEquality using (_≡_; refl)
+
+------------------------------------------------------------------------
+-- Examples
+
+-- Note that these examples are simple unit tests, because the type
+-- checker verifies them.
+
+private
+
+ module A = Monoid A
+
+ open import Data.Product.Categorical.Left A.rawMonoid b
+
+ _≈_ : Rel (A.Carrier × Lift a B) (e ⊔ a ⊔ b)
+ _≈_ = Pointwise A._≈_ _≡_
+
+ open RawFunctor functor
+
+ -- This type to the right of × needs to be a "lifted" version of (B : Set b)
+ -- that lives in the universe (Set (a ⊔ b)).
+ fmapIdₗ : (x : A.Carrier × Lift a B) → (id <$> x) ≈ x
+ fmapIdₗ x = A.refl , refl
+
+ open RawMonad monad
+
+ -- Now, let's show that "return" is a unit for >>=. We use Lift in exactly
+ -- the same way as above. The data (x : B) then needs to be "lifted" to
+ -- this new type (Lift B).
+ returnUnitL : ∀ {x : B} {f : Lift a B → A.Carrier × Lift a B} →
+ ((return (lift x)) >>= f) ≈ f (lift x)
+ returnUnitL = A.identityˡ _ , refl
+
+ returnUnitR : {x : A.Carrier × Lift a B} → (x >>= return) ≈ x
+ returnUnitR = A.identityʳ _ , refl
+
+ -- And another (limited version of a) monad law...
+ bindCompose : ∀ {f g : Lift a B → A.Carrier × Lift a B} →
+ {x : A.Carrier × Lift a B} →
+ ((x >>= f) >>= g) ≈ (x >>= (λ y → (f y >>= g)))
+ bindCompose = A.assoc _ _ _ , refl
diff --git a/src/Data/Product/Categorical/Left.agda b/src/Data/Product/Categorical/Left.agda
new file mode 100644
index 0000000..0375a14
--- /dev/null
+++ b/src/Data/Product/Categorical/Left.agda
@@ -0,0 +1,79 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Left-biased universe-sensitive functor and monad instances for the
+-- Product type.
+--
+-- To minimize the universe level of the RawFunctor, we require that
+-- elements of B are "lifted" to a copy of B at a higher universe level
+-- (a ⊔ b). See the Data.Product.Categorical.Examples for how this is
+-- done.
+------------------------------------------------------------------------
+
+open import Algebra
+open import Level
+
+module Data.Product.Categorical.Left
+ {a e} (A : RawMonoid a e) (b : Level) where
+
+open import Data.Product
+import Data.Product.Categorical.Left.Base as Base
+open import Category.Applicative using (RawApplicative)
+open import Category.Monad using (RawMonad; RawMonadT)
+open import Function using (id; flip; _∘_; _∘′_)
+import Function.Identity.Categorical as Id
+
+open RawMonoid A
+
+------------------------------------------------------------------------
+-- Re-export the base contents publically
+
+open Base Carrier b public
+
+------------------------------------------------------------------------
+-- Basic records
+
+applicative : RawApplicative Productₗ
+applicative = record
+ { pure = ε ,_
+ ; _⊛_ = zip _∙_ id
+ }
+
+-- The monad instance also requires some mucking about with universe levels.
+monadT : RawMonadT (_∘′ Productₗ)
+monadT M = record
+ { return = pure ∘′ (ε ,_)
+ ; _>>=_ = λ ma f → ma >>= uncurry λ a x → map₁ (a ∙_) <$> f x
+ } where open RawMonad M
+
+monad : RawMonad Productₗ
+monad = monadT Id.monad
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {F} (App : RawApplicative {a ⊔ b} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → Productₗ (F A) → F (Productₗ A)
+ sequenceA (x , fa) = (x ,_) <$> fa
+
+ mapA : ∀ {A B} → (A → F B) → Productₗ A → F (Productₗ B)
+ mapA f = sequenceA ∘ map₂ f
+
+ forA : ∀ {A B} → Productₗ A → (A → F B) → F (Productₗ B)
+ forA = flip mapA
+
+module _ {M} (Mon : RawMonad {a ⊔ b} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM : ∀ {A} → Productₗ (M A) → M (Productₗ A)
+ sequenceM = sequenceA App
+
+ mapM : ∀ {A B} → (A → M B) → Productₗ A → M (Productₗ B)
+ mapM = mapA App
+
+ forM : ∀ {A B} → Productₗ A → (A → M B) → M (Productₗ B)
+ forM = forA App
diff --git a/src/Data/Product/Categorical/Left/Base.agda b/src/Data/Product/Categorical/Left/Base.agda
new file mode 100644
index 0000000..5092b98
--- /dev/null
+++ b/src/Data/Product/Categorical/Left/Base.agda
@@ -0,0 +1,35 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Base definitions for the left-biased universe-sensitive functor and
+-- monad instances for the Product type.
+--
+-- To minimize the universe level of the RawFunctor, we require that
+-- elements of B are "lifted" to a copy of B at a higher universe level
+-- (a ⊔ b). See the Data.Product.Categorical.Examples for how this is
+-- done.
+------------------------------------------------------------------------
+
+open import Level
+
+module Data.Product.Categorical.Left.Base
+ {a} (A : Set a) (b : Level) where
+
+open import Data.Product using (_×_; map₂; proj₁; proj₂; <_,_>)
+open import Category.Functor using (RawFunctor)
+open import Category.Comonad using (RawComonad)
+
+------------------------------------------------------------------------
+-- Definitions
+
+Productₗ : Set (a ⊔ b) → Set (a ⊔ b)
+Productₗ B = A × B
+
+functor : RawFunctor Productₗ
+functor = record { _<$>_ = λ f → map₂ f }
+
+comonad : RawComonad Productₗ
+comonad = record
+ { extract = proj₂
+ ; extend = < proj₁ ,_>
+ }
diff --git a/src/Data/Product/Categorical/Right.agda b/src/Data/Product/Categorical/Right.agda
new file mode 100644
index 0000000..7011fef
--- /dev/null
+++ b/src/Data/Product/Categorical/Right.agda
@@ -0,0 +1,78 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Right-biased universe-sensitive functor and monad instances for the
+-- Product type.
+--
+-- To minimize the universe level of the RawFunctor, we require that
+-- elements of B are "lifted" to a copy of B at a higher universe level
+-- (a ⊔ b). See the Data.Product.Categorical.Examples for how this is
+-- done.
+------------------------------------------------------------------------
+
+open import Algebra
+open import Level
+
+module Data.Product.Categorical.Right
+ (a : Level) {b e} (B : RawMonoid b e) where
+
+open import Data.Product
+import Data.Product.Categorical.Right.Base as Base
+open import Category.Applicative using (RawApplicative)
+open import Category.Monad using (RawMonad; RawMonadT)
+open import Function using (id; flip; _∘_; _∘′_)
+import Function.Identity.Categorical as Id
+
+open RawMonoid B
+
+------------------------------------------------------------------------
+-- Re-export the base contents publically
+
+open Base Carrier a public
+
+------------------------------------------------------------------------
+-- Basic records
+
+applicative : RawApplicative Productᵣ
+applicative = record
+ { pure = _, ε
+ ; _⊛_ = zip id _∙_
+ }
+
+monadT : RawMonadT (_∘′ Productᵣ)
+monadT M = record
+ { return = pure ∘′ (_, ε)
+ ; _>>=_ = λ ma f → ma >>= uncurry λ x b → map₂ (b ∙_) <$> f x
+ } where open RawMonad M
+
+monad : RawMonad Productᵣ
+monad = monadT Id.monad
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {F} (App : RawApplicative {a ⊔ b} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → Productᵣ (F A) → F (Productᵣ A)
+ sequenceA (fa , y) = (_, y) <$> fa
+
+ mapA : ∀ {A B} → (A → F B) → Productᵣ A → F (Productᵣ B)
+ mapA f = sequenceA ∘ map₁ f
+
+ forA : ∀ {A B} → Productᵣ A → (A → F B) → F (Productᵣ B)
+ forA = flip mapA
+
+module _ {M} (Mon : RawMonad {a ⊔ b} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM : ∀ {A} → Productᵣ (M A) → M (Productᵣ A)
+ sequenceM = sequenceA App
+
+ mapM : ∀ {A B} → (A → M B) → Productᵣ A → M (Productᵣ B)
+ mapM = mapA App
+
+ forM : ∀ {A B} → Productᵣ A → (A → M B) → M (Productᵣ B)
+ forM = forA App
diff --git a/src/Data/Product/Categorical/Right/Base.agda b/src/Data/Product/Categorical/Right/Base.agda
new file mode 100644
index 0000000..af2d046
--- /dev/null
+++ b/src/Data/Product/Categorical/Right/Base.agda
@@ -0,0 +1,35 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Base definitions for the right-biased universe-sensitive functor
+-- and monad instances for the Product type.
+--
+-- To minimize the universe level of the RawFunctor, we require that
+-- elements of B are "lifted" to a copy of B at a higher universe level
+-- (a ⊔ b). See the Data.Product.Categorical.Examples for how this is
+-- done.
+------------------------------------------------------------------------
+
+open import Level
+
+module Data.Product.Categorical.Right.Base
+ {b} (B : Set b) (a : Level) where
+
+open import Data.Product using (_×_; map₁; proj₁; proj₂; <_,_>)
+open import Category.Functor using (RawFunctor)
+open import Category.Comonad using (RawComonad)
+
+------------------------------------------------------------------------
+-- Definitions
+
+Productᵣ : Set (a ⊔ b) → Set (a ⊔ b)
+Productᵣ A = A × B
+
+functor : RawFunctor Productᵣ
+functor = record { _<$>_ = map₁ }
+
+comonad : RawComonad Productᵣ
+comonad = record
+ { extract = proj₁
+ ; extend = <_, proj₂ >
+ }
diff --git a/src/Data/Product/N-ary.agda b/src/Data/Product/N-ary.agda
index 1bec530..746f10c 100644
--- a/src/Data/Product/N-ary.agda
+++ b/src/Data/Product/N-ary.agda
@@ -12,55 +12,134 @@
module Data.Product.N-ary where
-open import Data.Nat hiding (_^_)
-open import Data.Product
+open import Data.Nat as Nat hiding (_^_)
+open import Data.Fin hiding (lift)
+open import Data.Product as P using (_×_ ; _,_ ; ∃₂ ; uncurry)
+open import Data.Sum using (_⊎_)
open import Data.Unit
-open import Data.Vec
-open import Function.Inverse
+open import Data.Empty
+open import Function
open import Level using (Lift; lift)
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Relation.Binary.PropositionalEquality using (_≡_)
--- N-ary product.
+-- Types and patterns
+------------------------------------------------------------------------
-infix 8 _^_
+pattern 2+_ n = suc (suc n)
+infix 8 _^_
_^_ : ∀ {ℓ} → Set ℓ → ℕ → Set ℓ
-A ^ 0 = Lift ⊤
-A ^ 1 = A
-A ^ suc (suc n) = A × A ^ suc n
-
--- Conversions.
-
-↔Vec : ∀ {a} {A : Set a} n → A ^ n ↔ Vec A n
-↔Vec n = record
- { to = P.→-to-⟶ (toVec n)
- ; from = P.→-to-⟶ fromVec
- ; inverse-of = record
- { left-inverse-of = fromVec∘toVec n
- ; right-inverse-of = toVec∘fromVec
- }
- }
- where
- toVec : ∀ {a} {A : Set a} n → A ^ n → Vec A n
- toVec 0 (lift tt) = []
- toVec 1 x = [ x ]
- toVec (suc (suc n)) (x , xs) = x ∷ toVec _ xs
-
- fromVec : ∀ {a} {A : Set a} {n} → Vec A n → A ^ n
- fromVec [] = lift tt
- fromVec (x ∷ []) = x
- fromVec (x ∷ y ∷ xs) = (x , fromVec (y ∷ xs))
-
- fromVec∘toVec : ∀ {a} {A : Set a} n (xs : A ^ n) →
- fromVec (toVec n xs) ≡ xs
- fromVec∘toVec 0 (lift tt) = P.refl
- fromVec∘toVec 1 x = P.refl
- fromVec∘toVec 2 (x , y) = P.refl
- fromVec∘toVec (suc (suc (suc n))) (x , y , xs) =
- P.cong (_,_ x) (fromVec∘toVec (suc (suc n)) (y , xs))
-
- toVec∘fromVec : ∀ {a} {A : Set a} {n} (xs : Vec A n) →
- toVec n (fromVec xs) ≡ xs
- toVec∘fromVec [] = P.refl
- toVec∘fromVec (x ∷ []) = P.refl
- toVec∘fromVec (x ∷ y ∷ xs) = P.cong (_∷_ x) (toVec∘fromVec (y ∷ xs))
+A ^ 0 = Lift _ ⊤
+A ^ 1 = A
+A ^ 2+ n = A × A ^ suc n
+
+pattern [] = lift tt
+
+module _ {a} {A : Set a} where
+
+ infix 3 _∈[_]_
+ _∈[_]_ : A → ∀ n → A ^ n → Set a
+ a ∈[ 0 ] as = Lift _ ⊥
+ a ∈[ 1 ] a′ = a ≡ a′
+ a ∈[ 2+ n ] a′ , as = a ≡ a′ ⊎ a ∈[ suc n ] as
+
+-- Basic operations
+------------------------------------------------------------------------
+
+module _ {a} {A : Set a} where
+
+ cons : ∀ n → A → A ^ n → A ^ suc n
+ cons 0 a _ = a
+ cons (suc n) a as = a , as
+
+ uncons : ∀ n → A ^ suc n → A × A ^ n
+ uncons 0 a = a , lift tt
+ uncons (suc n) (a , as) = a , as
+
+ head : ∀ n → A ^ suc n → A
+ head n as = P.proj₁ (uncons n as)
+
+ tail : ∀ n → A ^ suc n → A ^ n
+ tail n as = P.proj₂ (uncons n as)
+
+ lookup : ∀ {n} (k : Fin n) → A ^ n → A
+ lookup {suc n} zero = head n
+ lookup {suc n} (suc k) = lookup k ∘′ tail n
+
+ replicate : ∀ n → A → A ^ n
+ replicate 0 a = []
+ replicate 1 a = a
+ replicate (2+ n) a = a , replicate (suc n) a
+
+ tabulate : ∀ n → (Fin n → A) → A ^ n
+ tabulate 0 f = []
+ tabulate 1 f = f zero
+ tabulate (2+ n) f = f zero , tabulate (suc n) (f ∘′ suc)
+
+ append : ∀ m n → A ^ m → A ^ n → A ^ (m Nat.+ n)
+ append 0 n xs ys = ys
+ append 1 n x ys = cons n x ys
+ append (2+ m) n (x , xs) ys = x , append (suc m) n xs ys
+
+ splitAt : ∀ m n → A ^ (m Nat.+ n) → A ^ m × A ^ n
+ splitAt 0 n xs = [] , xs
+ splitAt (suc m) n xs =
+ let (ys , zs) = splitAt m n (tail (m Nat.+ n) xs) in
+ cons m (head (m Nat.+ n) xs) ys , zs
+
+
+-- Manipulating N-ary products
+------------------------------------------------------------------------
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ map : (A → B) → ∀ n → A ^ n → B ^ n
+ map f 0 as = lift tt
+ map f 1 a = f a
+ map f (2+ n) (a , as) = f a , map f (suc n) as
+
+ ap : ∀ n → (A → B) ^ n → A ^ n → B ^ n
+ ap 0 fs ts = []
+ ap 1 f t = f t
+ ap (2+ n) (f , fs) (t , ts) = f t , ap (suc n) fs ts
+
+module _ {a p} {A : Set a} {P : ℕ → Set p} where
+
+ foldr : P 0 → (A → P 1) → (∀ n → A → P (suc n) → P (2+ n)) →
+ ∀ n → A ^ n → P n
+ foldr p0 p1 p2+ 0 as = p0
+ foldr p0 p1 p2+ 1 a = p1 a
+ foldr p0 p1 p2+ (2+ n) (a , as) = p2+ n a (foldr p0 p1 p2+ (suc n) as)
+
+foldl : ∀ {a p} {A : Set a} (P : ℕ → Set p) →
+ P 0 → (A → P 1) → (∀ n → A → P (suc n) → P (2+ n)) →
+ ∀ n → A ^ n → P n
+foldl P p0 p1 p2+ 0 as = p0
+foldl P p0 p1 p2+ 1 a = p1 a
+foldl P p0 p1 p2+ (2+ n) (a , as) = let p1′ = p1 a in
+ foldl (P ∘′ suc) p1′ (λ a → p2+ 0 a p1′) (p2+ ∘ suc) (suc n) as
+
+module _ {a} {A : Set a} where
+
+ reverse : ∀ n → A ^ n → A ^ n
+ reverse = foldl (A ^_) [] id (λ n → _,_)
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
+
+ zipWith : (A → B → C) → ∀ n → A ^ n → B ^ n → C ^ n
+ zipWith f 0 as bs = []
+ zipWith f 1 a b = f a b
+ zipWith f (2+ n) (a , as) (b , bs) = f a b , zipWith f (suc n) as bs
+
+ unzipWith : (A → B × C) → ∀ n → A ^ n → B ^ n × C ^ n
+ unzipWith f 0 as = [] , []
+ unzipWith f 1 a = f a
+ unzipWith f (2+ n) (a , as) = P.zip _,_ _,_ (f a) (unzipWith f (suc n) as)
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ zip : ∀ n → A ^ n → B ^ n → (A × B) ^ n
+ zip = zipWith _,_
+
+ unzip : ∀ n → (A × B) ^ n → A ^ n × B ^ n
+ unzip = unzipWith id
diff --git a/src/Data/Product/N-ary/Categorical.agda b/src/Data/Product/N-ary/Categorical.agda
new file mode 100644
index 0000000..03849df
--- /dev/null
+++ b/src/Data/Product/N-ary/Categorical.agda
@@ -0,0 +1,59 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A categorical view of N-ary products
+------------------------------------------------------------------------
+
+module Data.Product.N-ary.Categorical where
+
+open import Agda.Builtin.Nat
+open import Data.Product hiding (map)
+open import Data.Product.N-ary
+open import Function
+
+open import Category.Functor
+open import Category.Applicative
+open import Category.Monad
+
+------------------------------------------------------------------------
+-- Functor and applicative
+
+functor : ∀ {ℓ} n → RawFunctor {ℓ} (_^ n)
+functor n = record { _<$>_ = λ f → map f n }
+
+applicative : ∀ {ℓ} n → RawApplicative {ℓ} (_^ n)
+applicative n = record
+ { pure = replicate n
+ ; _⊛_ = ap n
+ }
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {f F} (App : RawApplicative {f} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {n A} → F A ^ n → F (A ^ n)
+ sequenceA {0} _ = pure _
+ sequenceA {1} fa = fa
+ sequenceA {2+ n} (fa , fas) = _,_ <$> fa ⊛ sequenceA fas
+
+ mapA : ∀ {n a} {A : Set a} {B} → (A → F B) → A ^ n → F (B ^ n)
+ mapA f = sequenceA ∘ map f _
+
+ forA : ∀ {n a} {A : Set a} {B} → A ^ n → (A → F B) → F (B ^ n)
+ forA = flip mapA
+
+module _ {m M} (Mon : RawMonad {m} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM : ∀ {n A} → M A ^ n → M (A ^ n)
+ sequenceM = sequenceA App
+
+ mapM : ∀ {n a} {A : Set a} {B} → (A → M B) → A ^ n → M (B ^ n)
+ mapM = mapA App
+
+ forM : ∀ {n a} {A : Set a} {B} → A ^ n → (A → M B) → M (B ^ n)
+ forM = forA App
diff --git a/src/Data/Product/N-ary/Properties.agda b/src/Data/Product/N-ary/Properties.agda
new file mode 100644
index 0000000..8c3aa3c
--- /dev/null
+++ b/src/Data/Product/N-ary/Properties.agda
@@ -0,0 +1,89 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of n-ary products
+------------------------------------------------------------------------
+
+module Data.Product.N-ary.Properties where
+
+open import Data.Nat.Base hiding (_^_)
+open import Data.Product
+open import Data.Product.N-ary
+open import Data.Vec using (Vec; _∷_)
+open import Function.Inverse using (_↔_; inverse)
+open import Relation.Binary.PropositionalEquality as P
+open ≡-Reasoning
+
+------------------------------------------------------------------------
+-- Basic proofs
+
+module _ {a} {A : Set a} where
+
+ cons-head-tail-identity : ∀ n (as : A ^ suc n) → cons n (head n as) (tail n as) ≡ as
+ cons-head-tail-identity 0 as = P.refl
+ cons-head-tail-identity (suc n) as = P.refl
+
+ head-cons-identity : ∀ n a (as : A ^ n) → head n (cons n a as) ≡ a
+ head-cons-identity 0 a as = P.refl
+ head-cons-identity (suc n) a as = P.refl
+
+ tail-cons-identity : ∀ n a (as : A ^ n) → tail n (cons n a as) ≡ as
+ tail-cons-identity 0 a as = P.refl
+ tail-cons-identity (suc n) a as = P.refl
+
+ append-cons-commute : ∀ m n a (xs : A ^ m) ys →
+ append (suc m) n (cons m a xs) ys ≡ cons (m + n) a (append m n xs ys)
+ append-cons-commute 0 n a xs ys = P.refl
+ append-cons-commute (suc m) n a xs ys = P.refl
+
+ append-splitAt-identity : ∀ m n (as : A ^ (m + n)) → uncurry (append m n) (splitAt m n as) ≡ as
+ append-splitAt-identity 0 n as = P.refl
+ append-splitAt-identity (suc m) n as = begin
+ let x = head (m + n) as in
+ let (xs , ys) = splitAt m n (tail (m + n) as) in
+ append (suc m) n (cons m (head (m + n) as) xs) ys
+ ≡⟨ append-cons-commute m n x xs ys ⟩
+ cons (m + n) x (append m n xs ys)
+ ≡⟨ P.cong (cons (m + n) x) (append-splitAt-identity m n (tail (m + n) as)) ⟩
+ cons (m + n) x (tail (m + n) as)
+ ≡⟨ cons-head-tail-identity (m + n) as ⟩
+ as
+ ∎
+
+------------------------------------------------------------------------
+-- Conversion to and from Vec
+
+module _ {a} {A : Set a} where
+
+ toVec : ∀ n → A ^ n → Vec A n
+ toVec 0 _ = Vec.[]
+ toVec (suc n) xs = head n xs Vec.∷ toVec n (tail n xs)
+
+ fromVec : ∀ {n} → Vec A n → A ^ n
+ fromVec Vec.[] = []
+ fromVec {suc n} (x Vec.∷ xs) = cons n x (fromVec xs)
+
+ fromVec∘toVec : ∀ n (xs : A ^ n) → fromVec (toVec n xs) ≡ xs
+ fromVec∘toVec 0 _ = P.refl
+ fromVec∘toVec (suc n) xs = begin
+ cons n (head n xs) (fromVec (toVec n (tail n xs)))
+ ≡⟨ P.cong (cons n (head n xs)) (fromVec∘toVec n (tail n xs)) ⟩
+ cons n (head n xs) (tail n xs)
+ ≡⟨ cons-head-tail-identity n xs ⟩
+ xs ∎
+
+ toVec∘fromVec : ∀ {n} (xs : Vec A n) → toVec n (fromVec xs) ≡ xs
+ toVec∘fromVec Vec.[] = P.refl
+ toVec∘fromVec {suc n} (x Vec.∷ xs) = begin
+ head n (cons n x (fromVec xs)) Vec.∷ toVec n (tail n (cons n x (fromVec xs)))
+ ≡⟨ P.cong₂ (λ x xs → x Vec.∷ toVec n xs) hd-prf tl-prf ⟩
+ x Vec.∷ toVec n (fromVec xs)
+ ≡⟨ P.cong (x Vec.∷_) (toVec∘fromVec xs) ⟩
+ x Vec.∷ xs
+ ∎ where
+
+ hd-prf = head-cons-identity n x (fromVec xs)
+ tl-prf = tail-cons-identity n x (fromVec xs)
+
+ ↔Vec : ∀ n → A ^ n ↔ Vec A n
+ ↔Vec n = inverse (toVec n) fromVec (fromVec∘toVec n) toVec∘fromVec
diff --git a/src/Data/Product/Properties.agda b/src/Data/Product/Properties.agda
new file mode 100644
index 0000000..a7831b9
--- /dev/null
+++ b/src/Data/Product/Properties.agda
@@ -0,0 +1,19 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of products
+------------------------------------------------------------------------
+
+module Data.Product.Properties where
+
+open import Data.Product
+open import Function
+open import Relation.Binary.PropositionalEquality
+
+module _ {a b} {A : Set a} {B : A → Set b} where
+
+ ,-injectiveˡ : ∀ {a c} {b : B a} {d : B c} → (a , b) ≡ (c , d) → a ≡ c
+ ,-injectiveˡ refl = refl
+
+ ,-injectiveʳ : ∀ {a} {b c : B a} → (Σ A B ∋ (a , b)) ≡ (a , c) → b ≡ c
+ ,-injectiveʳ refl = refl
diff --git a/src/Data/Product/Relation/Lex/NonStrict.agda b/src/Data/Product/Relation/Lex/NonStrict.agda
new file mode 100644
index 0000000..6de292c
--- /dev/null
+++ b/src/Data/Product/Relation/Lex/NonStrict.agda
@@ -0,0 +1,213 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Lexicographic products of binary relations
+------------------------------------------------------------------------
+
+-- The definition of lexicographic product used here is suitable if
+-- the left-hand relation is a (non-strict) partial order.
+
+module Data.Product.Relation.Lex.NonStrict where
+
+open import Data.Product using (_×_; _,_; proj₁; proj₂)
+open import Data.Sum using (inj₁; inj₂)
+open import Relation.Binary
+open import Relation.Binary.Consequences
+import Relation.Binary.Construct.NonStrictToStrict as Conv
+open import Data.Product.Relation.Pointwise.NonDependent as Pointwise
+ using (Pointwise)
+import Data.Product.Relation.Lex.Strict as Strict
+
+module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where
+
+------------------------------------------------------------------------
+-- A lexicographic ordering over products
+
+ ×-Lex : (_≈₁_ _≤₁_ : Rel A₁ ℓ₁) (_≤₂_ : Rel A₂ ℓ₂) → Rel (A₁ × A₂) _
+ ×-Lex _≈₁_ _≤₁_ _≤₂_ = Strict.×-Lex _≈₁_ (Conv._<_ _≈₁_ _≤₁_) _≤₂_
+
+------------------------------------------------------------------------
+-- Some properties which are preserved by ×-Lex (under certain
+-- assumptions).
+
+ ×-reflexive : ∀ _≈₁_ _≤₁_ {_≈₂_} _≤₂_ →
+ _≈₂_ ⇒ _≤₂_ →
+ (Pointwise _≈₁_ _≈₂_) ⇒ (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-reflexive _≈₁_ _≤₁_ _≤₂_ refl₂ =
+ Strict.×-reflexive _≈₁_ (Conv._<_ _≈₁_ _≤₁_) _≤₂_ refl₂
+
+ ×-transitive : ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ →
+ ∀ {_≤₂_} → Transitive _≤₂_ →
+ Transitive (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-transitive {_≈₁_} {_≤₁_} po₁ {_≤₂_} trans₂ =
+ Strict.×-transitive
+ {_<₁_ = Conv._<_ _≈₁_ _≤₁_}
+ isEquivalence (Conv.<-resp-≈ _ _ isEquivalence ≤-resp-≈)
+ (Conv.<-trans _ _ po₁)
+ {_≤₂_} trans₂
+ where open IsPartialOrder po₁
+
+ ×-antisymmetric :
+ ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ →
+ ∀ {_≈₂_ _≤₂_} → Antisymmetric _≈₂_ _≤₂_ →
+ Antisymmetric (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-antisymmetric {_≈₁_} {_≤₁_}
+ po₁ {_≤₂_ = _≤₂_} antisym₂ =
+ Strict.×-antisymmetric {_<₁_ = Conv._<_ _≈₁_ _≤₁_}
+ ≈-sym₁ irrefl₁ asym₁
+ {_≤₂_ = _≤₂_} antisym₂
+ where
+ open IsPartialOrder po₁
+ open Eq renaming (refl to ≈-refl₁; sym to ≈-sym₁)
+
+ irrefl₁ : Irreflexive _≈₁_ (Conv._<_ _≈₁_ _≤₁_)
+ irrefl₁ = Conv.<-irrefl _≈₁_ _≤₁_
+
+ asym₁ : Asymmetric (Conv._<_ _≈₁_ _≤₁_)
+ asym₁ = trans∧irr⟶asym {_≈_ = _≈₁_}
+ ≈-refl₁ (Conv.<-trans _ _ po₁) irrefl₁
+
+ ×-respects₂ :
+ ∀ {_≈₁_ _≤₁_} → IsEquivalence _≈₁_ → _≤₁_ Respects₂ _≈₁_ →
+ ∀ {_≈₂_ _≤₂_ : Rel A₂ ℓ₂} → _≤₂_ Respects₂ _≈₂_ →
+ (×-Lex _≈₁_ _≤₁_ _≤₂_) Respects₂ (Pointwise _≈₁_ _≈₂_)
+ ×-respects₂ eq₁ resp₁ resp₂ =
+ Strict.×-respects₂ eq₁ (Conv.<-resp-≈ _ _ eq₁ resp₁) resp₂
+
+ ×-decidable : ∀ {_≈₁_ _≤₁_} → Decidable _≈₁_ → Decidable _≤₁_ →
+ ∀ {_≤₂_} → Decidable _≤₂_ →
+ Decidable (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-decidable dec-≈₁ dec-≤₁ dec-≤₂ =
+ Strict.×-decidable dec-≈₁ (Conv.<-decidable _ _ dec-≈₁ dec-≤₁)
+ dec-≤₂
+
+ ×-total : ∀ {_≈₁_ _≤₁_} → Symmetric _≈₁_ → Decidable _≈₁_ →
+ Antisymmetric _≈₁_ _≤₁_ → Total _≤₁_ →
+ ∀ {_≤₂_} → Total _≤₂_ →
+ Total (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-total {_≈₁_} {_≤₁_} sym₁ dec₁ antisym₁ total₁ {_≤₂_} total₂ =
+ total
+ where
+ tri₁ : Trichotomous _≈₁_ (Conv._<_ _≈₁_ _≤₁_)
+ tri₁ = Conv.<-trichotomous _ _ sym₁ dec₁ antisym₁ total₁
+
+ total : Total (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ total x y with tri₁ (proj₁ x) (proj₁ y)
+ ... | tri< x₁<y₁ x₁≉y₁ x₁≯y₁ = inj₁ (inj₁ x₁<y₁)
+ ... | tri> x₁≮y₁ x₁≉y₁ x₁>y₁ = inj₂ (inj₁ x₁>y₁)
+ ... | tri≈ x₁≮y₁ x₁≈y₁ x₁≯y₁ with total₂ (proj₂ x) (proj₂ y)
+ ... | inj₁ x₂≤y₂ = inj₁ (inj₂ (x₁≈y₁ , x₂≤y₂))
+ ... | inj₂ x₂≥y₂ = inj₂ (inj₂ (sym₁ x₁≈y₁ , x₂≥y₂))
+
+ -- Some collections of properties which are preserved by ×-Lex
+ -- (under certain assumptions).
+
+ ×-isPartialOrder :
+ ∀ {_≈₁_ _≤₁_} → IsPartialOrder _≈₁_ _≤₁_ →
+ ∀ {_≈₂_ _≤₂_} → IsPartialOrder _≈₂_ _≤₂_ →
+ IsPartialOrder (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-isPartialOrder {_≈₁_} {_≤₁_} po₁
+ {_≤₂_ = _≤₂_} po₂ = record
+ { isPreorder = record
+ { isEquivalence = Pointwise.×-isEquivalence
+ (isEquivalence po₁)
+ (isEquivalence po₂)
+ ; reflexive = ×-reflexive _≈₁_ _≤₁_ _≤₂_ (reflexive po₂)
+ ; trans = ×-transitive po₁ {_≤₂_ = _≤₂_} (trans po₂)
+ }
+ ; antisym = ×-antisymmetric {_≤₁_ = _≤₁_} po₁
+ {_≤₂_ = _≤₂_} (antisym po₂)
+ }
+ where open IsPartialOrder
+
+ ×-isTotalOrder :
+ ∀ {_≈₁_ _≤₁_} → Decidable _≈₁_ → IsTotalOrder _≈₁_ _≤₁_ →
+ ∀ {_≈₂_ _≤₂_} → IsTotalOrder _≈₂_ _≤₂_ →
+ IsTotalOrder (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-isTotalOrder {_≤₁_ = _≤₁_} ≈₁-dec to₁ {_≤₂_ = _≤₂_} to₂ = record
+ { isPartialOrder = ×-isPartialOrder
+ (isPartialOrder to₁) (isPartialOrder to₂)
+ ; total = ×-total {_≤₁_ = _≤₁_} (Eq.sym to₁) ≈₁-dec
+ (antisym to₁) (total to₁)
+ {_≤₂_ = _≤₂_} (total to₂)
+ }
+ where open IsTotalOrder
+
+ ×-isDecTotalOrder :
+ ∀ {_≈₁_ _≤₁_} → IsDecTotalOrder _≈₁_ _≤₁_ →
+ ∀ {_≈₂_ _≤₂_} → IsDecTotalOrder _≈₂_ _≤₂_ →
+ IsDecTotalOrder (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _≤₁_ _≤₂_)
+ ×-isDecTotalOrder {_≤₁_ = _≤₁_} to₁ {_≤₂_ = _≤₂_} to₂ = record
+ { isTotalOrder = ×-isTotalOrder (_≟_ to₁)
+ (isTotalOrder to₁)
+ (isTotalOrder to₂)
+ ; _≟_ = Pointwise.×-decidable (_≟_ to₁) (_≟_ to₂)
+ ; _≤?_ = ×-decidable (_≟_ to₁) (_≤?_ to₁) (_≤?_ to₂)
+ }
+ where open IsDecTotalOrder
+
+------------------------------------------------------------------------
+-- "Packages" can also be combined.
+
+module _ {ℓ₁ ℓ₂ ℓ₃ ℓ₄} where
+
+ ×-poset : Poset ℓ₁ ℓ₂ _ → Poset ℓ₃ ℓ₄ _ → Poset _ _ _
+ ×-poset p₁ p₂ = record
+ { isPartialOrder = ×-isPartialOrder
+ (isPartialOrder p₁) (isPartialOrder p₂)
+ } where open Poset
+
+ ×-totalOrder : DecTotalOrder ℓ₁ ℓ₂ _ → TotalOrder ℓ₃ ℓ₄ _ →
+ TotalOrder _ _ _
+ ×-totalOrder t₁ t₂ = record
+ { isTotalOrder = ×-isTotalOrder T₁._≟_ T₁.isTotalOrder T₂.isTotalOrder
+ }
+ where
+ module T₁ = DecTotalOrder t₁
+ module T₂ = TotalOrder t₂
+
+ ×-decTotalOrder : DecTotalOrder ℓ₁ ℓ₂ _ → DecTotalOrder ℓ₃ ℓ₄ _ →
+ DecTotalOrder _ _ _
+ ×-decTotalOrder t₁ t₂ = record
+ { isDecTotalOrder = ×-isDecTotalOrder
+ (isDecTotalOrder t₁) (isDecTotalOrder t₂)
+ } where open DecTotalOrder
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+_×-isPartialOrder_ = ×-isPartialOrder
+{-# WARNING_ON_USAGE _×-isPartialOrder_
+"Warning: _×-isPartialOrder_ was deprecated in v0.15.
+Please use ×-isPartialOrder instead."
+#-}
+_×-isDecTotalOrder_ = ×-isDecTotalOrder
+{-# WARNING_ON_USAGE _×-isDecTotalOrder_
+"Warning: _×-isDecTotalOrder_ was deprecated in v0.15.
+Please use ×-isDecTotalOrder instead."
+#-}
+_×-poset_ = ×-poset
+{-# WARNING_ON_USAGE _×-poset_
+"Warning: _×-poset_ was deprecated in v0.15.
+Please use ×-poset instead."
+#-}
+_×-totalOrder_ = ×-totalOrder
+{-# WARNING_ON_USAGE _×-totalOrder_
+"Warning: _×-totalOrder_ was deprecated in v0.15.
+Please use ×-totalOrder instead."
+#-}
+_×-decTotalOrder_ = ×-decTotalOrder
+{-# WARNING_ON_USAGE _×-decTotalOrder_
+"Warning: _×-decTotalOrder_ was deprecated in v0.15.
+Please use ×-decTotalOrder instead."
+#-}
+×-≈-respects₂ = ×-respects₂
+{-# WARNING_ON_USAGE ×-≈-respects₂
+"Warning: ×-≈-respects₂ was deprecated in v0.15.
+Please use ×-respects₂ instead."
+#-}
diff --git a/src/Data/Product/Relation/Lex/Strict.agda b/src/Data/Product/Relation/Lex/Strict.agda
new file mode 100644
index 0000000..b48fcac
--- /dev/null
+++ b/src/Data/Product/Relation/Lex/Strict.agda
@@ -0,0 +1,302 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Lexicographic products of binary relations
+------------------------------------------------------------------------
+
+-- The definition of lexicographic product used here is suitable if
+-- the left-hand relation is a strict partial order.
+
+module Data.Product.Relation.Lex.Strict where
+
+open import Data.Product
+open import Data.Product.Relation.Pointwise.NonDependent as Pointwise
+ using (Pointwise)
+open import Data.Sum using (inj₁; inj₂; _-⊎-_; [_,_])
+open import Data.Empty
+open import Function
+open import Level
+open import Relation.Nullary
+open import Relation.Nullary.Product
+open import Relation.Nullary.Sum
+open import Relation.Binary
+open import Relation.Binary.Consequences
+
+module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where
+
+------------------------------------------------------------------------
+-- A lexicographic ordering over products
+
+ ×-Lex : (_≈₁_ _<₁_ : Rel A₁ ℓ₁) → (_≤₂_ : Rel A₂ ℓ₂) → Rel (A₁ × A₂) _
+ ×-Lex _≈₁_ _<₁_ _≤₂_ =
+ (_<₁_ on proj₁) -⊎- (_≈₁_ on proj₁) -×- (_≤₂_ on proj₂)
+
+------------------------------------------------------------------------
+-- Some properties which are preserved by ×-Lex (under certain
+-- assumptions).
+
+ ×-reflexive : ∀ _≈₁_ _∼₁_ {_≈₂_ : Rel A₂ ℓ₂} _≤₂_ →
+ _≈₂_ ⇒ _≤₂_ → (Pointwise _≈₁_ _≈₂_) ⇒ (×-Lex _≈₁_ _∼₁_ _≤₂_)
+ ×-reflexive _ _ _ refl₂ = λ x≈y →
+ inj₂ (proj₁ x≈y , refl₂ (proj₂ x≈y))
+
+ ×-irreflexive : ∀ {_≈₁_ _<₁_} → Irreflexive _≈₁_ _<₁_ →
+ ∀ {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → Irreflexive _≈₂_ _<₂_ →
+ Irreflexive (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ ×-irreflexive ir₁ ir₂ x≈y (inj₁ x₁<y₁) = ir₁ (proj₁ x≈y) x₁<y₁
+ ×-irreflexive ir₁ ir₂ x≈y (inj₂ x≈<y) =
+ ir₂ (proj₂ x≈y) (proj₂ x≈<y)
+
+ ×-transitive : ∀ {_≈₁_ _<₁_} →
+ IsEquivalence _≈₁_ → _<₁_ Respects₂ _≈₁_ → Transitive _<₁_ →
+ ∀ {_≤₂_} → Transitive _≤₂_ → Transitive (×-Lex _≈₁_ _<₁_ _≤₂_)
+ ×-transitive {_≈₁_} {_<₁_} eq₁ resp₁ trans₁ {_≤₂_} trans₂ = trans
+ where
+ module Eq₁ = IsEquivalence eq₁
+
+ trans : Transitive (×-Lex _≈₁_ _<₁_ _≤₂_)
+ trans (inj₁ x₁<y₁) (inj₁ y₁<z₁) = inj₁ (trans₁ x₁<y₁ y₁<z₁)
+ trans (inj₁ x₁<y₁) (inj₂ y≈≤z) =
+ inj₁ (proj₁ resp₁ (proj₁ y≈≤z) x₁<y₁)
+ trans (inj₂ x≈≤y) (inj₁ y₁<z₁) =
+ inj₁ (proj₂ resp₁ (Eq₁.sym $ proj₁ x≈≤y) y₁<z₁)
+ trans (inj₂ x≈≤y) (inj₂ y≈≤z) =
+ inj₂ ( Eq₁.trans (proj₁ x≈≤y) (proj₁ y≈≤z)
+ , trans₂ (proj₂ x≈≤y) (proj₂ y≈≤z) )
+
+ ×-antisymmetric : ∀ {_≈₁_ _<₁_} →
+ Symmetric _≈₁_ → Irreflexive _≈₁_ _<₁_ → Asymmetric _<₁_ →
+ ∀ {_≈₂_ _≤₂_ : Rel A₂ ℓ₂} → Antisymmetric _≈₂_ _≤₂_ →
+ Antisymmetric (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _<₁_ _≤₂_)
+ ×-antisymmetric {_≈₁_} {_<₁_} sym₁ irrefl₁ asym₁
+ {_≈₂_} {_≤₂_} antisym₂ = antisym
+ where
+ antisym : Antisymmetric (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _<₁_ _≤₂_)
+ antisym (inj₁ x₁<y₁) (inj₁ y₁<x₁) =
+ ⊥-elim $ asym₁ x₁<y₁ y₁<x₁
+ antisym (inj₁ x₁<y₁) (inj₂ y≈≤x) =
+ ⊥-elim $ irrefl₁ (sym₁ $ proj₁ y≈≤x) x₁<y₁
+ antisym (inj₂ x≈≤y) (inj₁ y₁<x₁) =
+ ⊥-elim $ irrefl₁ (sym₁ $ proj₁ x≈≤y) y₁<x₁
+ antisym (inj₂ x≈≤y) (inj₂ y≈≤x) =
+ proj₁ x≈≤y , antisym₂ (proj₂ x≈≤y) (proj₂ y≈≤x)
+
+ ×-asymmetric : ∀ {_≈₁_ _<₁_} →
+ Symmetric _≈₁_ → _<₁_ Respects₂ _≈₁_ → Asymmetric _<₁_ →
+ ∀ {_<₂_} → Asymmetric _<₂_ →
+ Asymmetric (×-Lex _≈₁_ _<₁_ _<₂_)
+ ×-asymmetric {_≈₁_} {_<₁_} sym₁ resp₁ asym₁ {_<₂_} asym₂ = asym
+ where
+ irrefl₁ : Irreflexive _≈₁_ _<₁_
+ irrefl₁ = asym⟶irr resp₁ sym₁ asym₁
+
+ asym : Asymmetric (×-Lex _≈₁_ _<₁_ _<₂_)
+ asym (inj₁ x₁<y₁) (inj₁ y₁<x₁) = asym₁ x₁<y₁ y₁<x₁
+ asym (inj₁ x₁<y₁) (inj₂ y≈<x) = irrefl₁ (sym₁ $ proj₁ y≈<x) x₁<y₁
+ asym (inj₂ x≈<y) (inj₁ y₁<x₁) = irrefl₁ (sym₁ $ proj₁ x≈<y) y₁<x₁
+ asym (inj₂ x≈<y) (inj₂ y≈<x) = asym₂ (proj₂ x≈<y) (proj₂ y≈<x)
+
+ ×-respects₂ :
+ ∀ {_≈₁_ _<₁_} → IsEquivalence _≈₁_ → _<₁_ Respects₂ _≈₁_ →
+ {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → _<₂_ Respects₂ _≈₂_ →
+ (×-Lex _≈₁_ _<₁_ _<₂_) Respects₂ (Pointwise _≈₁_ _≈₂_)
+ ×-respects₂ {_≈₁_} {_<₁_} eq₁ resp₁
+ {_≈₂_} {_<₂_} resp₂ = resp¹ , resp²
+ where
+ _<_ = ×-Lex _≈₁_ _<₁_ _<₂_
+
+ open IsEquivalence eq₁ renaming (sym to sym₁; trans to trans₁)
+
+ resp¹ : ∀ {x} → (x <_) Respects (Pointwise _≈₁_ _≈₂_)
+ resp¹ y≈y' (inj₁ x₁<y₁) = inj₁ (proj₁ resp₁ (proj₁ y≈y') x₁<y₁)
+ resp¹ y≈y' (inj₂ x≈<y) =
+ inj₂ ( trans₁ (proj₁ x≈<y) (proj₁ y≈y')
+ , proj₁ resp₂ (proj₂ y≈y') (proj₂ x≈<y) )
+
+ resp² : ∀ {y} → (flip _<_ y) Respects (Pointwise _≈₁_ _≈₂_)
+ resp² x≈x' (inj₁ x₁<y₁) = inj₁ (proj₂ resp₁ (proj₁ x≈x') x₁<y₁)
+ resp² x≈x' (inj₂ x≈<y) =
+ inj₂ ( trans₁ (sym₁ $ proj₁ x≈x') (proj₁ x≈<y)
+ , proj₂ resp₂ (proj₂ x≈x') (proj₂ x≈<y) )
+
+ ×-decidable : ∀ {_≈₁_ _<₁_} → Decidable _≈₁_ → Decidable _<₁_ →
+ ∀ {_≤₂_} → Decidable _≤₂_ →
+ Decidable (×-Lex _≈₁_ _<₁_ _≤₂_)
+ ×-decidable dec-≈₁ dec-<₁ dec-≤₂ x y =
+ dec-<₁ (proj₁ x) (proj₁ y)
+ ⊎-dec
+ (dec-≈₁ (proj₁ x) (proj₁ y)
+ ×-dec
+ dec-≤₂ (proj₂ x) (proj₂ y))
+
+ ×-total₁ : ∀ {_≈₁_ _<₁_} → Total _<₁_ →
+ ∀ {_≤₂_} → Total (×-Lex _≈₁_ _<₁_ _≤₂_)
+ ×-total₁ total₁ x y with total₁ (proj₁ x) (proj₁ y)
+ ... | inj₁ x₁<y₁ = inj₁ (inj₁ x₁<y₁)
+ ... | inj₂ x₁>y₁ = inj₂ (inj₁ x₁>y₁)
+
+ ×-total₂ : ∀ {_≈₁_ _<₁_} → Symmetric _≈₁_ → Trichotomous _≈₁_ _<₁_ →
+ ∀ {_≤₂_} → Total _≤₂_ →
+ Total (×-Lex _≈₁_ _<₁_ _≤₂_)
+ ×-total₂ sym tri₁ total₂ x y with tri₁ (proj₁ x) (proj₁ y)
+ ... | tri< x₁<y₁ _ _ = inj₁ (inj₁ x₁<y₁)
+ ... | tri> _ _ y₁<x₁ = inj₂ (inj₁ y₁<x₁)
+ ... | tri≈ _ x₁≈y₁ _ with total₂ (proj₂ x) (proj₂ y)
+ ... | inj₁ x₂≤y₂ = inj₁ (inj₂ (x₁≈y₁ , x₂≤y₂))
+ ... | inj₂ y₂≤x₂ = inj₂ (inj₂ (sym x₁≈y₁ , y₂≤x₂))
+
+ ×-compare :
+ {_≈₁_ _<₁_ : Rel A₁ ℓ₁} → Symmetric _≈₁_ → Trichotomous _≈₁_ _<₁_ →
+ {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → Trichotomous _≈₂_ _<₂_ →
+ Trichotomous (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ ×-compare sym₁ cmp₁ cmp₂ (x₁ , x₂) (y₁ , y₂) with cmp₁ x₁ y₁
+ ... | (tri< x₁<y₁ x₁≉y₁ x₁≯y₁) =
+ tri< (inj₁ x₁<y₁)
+ (x₁≉y₁ ∘ proj₁)
+ [ x₁≯y₁ , x₁≉y₁ ∘ sym₁ ∘ proj₁ ]
+ ... | (tri> x₁≮y₁ x₁≉y₁ x₁>y₁) =
+ tri> [ x₁≮y₁ , x₁≉y₁ ∘ proj₁ ]
+ (x₁≉y₁ ∘ proj₁)
+ (inj₁ x₁>y₁)
+ ... | (tri≈ x₁≮y₁ x₁≈y₁ x₁≯y₁) with cmp₂ x₂ y₂
+ ... | (tri< x₂<y₂ x₂≉y₂ x₂≯y₂) =
+ tri< (inj₂ (x₁≈y₁ , x₂<y₂))
+ (x₂≉y₂ ∘ proj₂)
+ [ x₁≯y₁ , x₂≯y₂ ∘ proj₂ ]
+ ... | (tri> x₂≮y₂ x₂≉y₂ x₂>y₂) =
+ tri> [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ]
+ (x₂≉y₂ ∘ proj₂)
+ (inj₂ (sym₁ x₁≈y₁ , x₂>y₂))
+ ... | (tri≈ x₂≮y₂ x₂≈y₂ x₂≯y₂) =
+ tri≈ [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ]
+ (x₁≈y₁ , x₂≈y₂)
+ [ x₁≯y₁ , x₂≯y₂ ∘ proj₂ ]
+
+------------------------------------------------------------------------
+-- Collections of properties which are preserved by ×-Lex.
+
+ ×-isPreorder : ∀ {_≈₁_ _∼₁_} → IsPreorder _≈₁_ _∼₁_ →
+ ∀ {_≈₂_ _∼₂_} → IsPreorder _≈₂_ _∼₂_ →
+ IsPreorder (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _∼₁_ _∼₂_)
+ ×-isPreorder {_≈₁_} {_∼₁_} pre₁ {_∼₂_ = _∼₂_} pre₂ =
+ record
+ { isEquivalence = Pointwise.×-isEquivalence
+ (isEquivalence pre₁) (isEquivalence pre₂)
+ ; reflexive = ×-reflexive _≈₁_ _∼₁_ _∼₂_ (reflexive pre₂)
+ ; trans = ×-transitive
+ (isEquivalence pre₁) (∼-resp-≈ pre₁)
+ (trans pre₁) {_≤₂_ = _∼₂_} (trans pre₂)
+ }
+ where open IsPreorder
+
+ ×-isStrictPartialOrder :
+ ∀ {_≈₁_ _<₁_} → IsStrictPartialOrder _≈₁_ _<₁_ →
+ ∀ {_≈₂_ _<₂_} → IsStrictPartialOrder _≈₂_ _<₂_ →
+ IsStrictPartialOrder (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ ×-isStrictPartialOrder {_<₁_ = _<₁_} spo₁ {_<₂_ = _<₂_} spo₂ =
+ record
+ { isEquivalence = Pointwise.×-isEquivalence
+ (isEquivalence spo₁) (isEquivalence spo₂)
+ ; irrefl = ×-irreflexive {_<₁_ = _<₁_} (irrefl spo₁)
+ {_<₂_ = _<₂_} (irrefl spo₂)
+ ; trans = ×-transitive
+ {_<₁_ = _<₁_} (isEquivalence spo₁)
+ (<-resp-≈ spo₁) (trans spo₁)
+ {_≤₂_ = _<₂_} (trans spo₂)
+ ; <-resp-≈ = ×-respects₂ (isEquivalence spo₁)
+ (<-resp-≈ spo₁)
+ (<-resp-≈ spo₂)
+ }
+ where open IsStrictPartialOrder
+
+ ×-isStrictTotalOrder :
+ ∀ {_≈₁_ _<₁_} → IsStrictTotalOrder _≈₁_ _<₁_ →
+ ∀ {_≈₂_ _<₂_} → IsStrictTotalOrder _≈₂_ _<₂_ →
+ IsStrictTotalOrder (Pointwise _≈₁_ _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_)
+ ×-isStrictTotalOrder {_<₁_ = _<₁_} spo₁ {_<₂_ = _<₂_} spo₂ =
+ record
+ { isEquivalence = Pointwise.×-isEquivalence
+ (isEquivalence spo₁) (isEquivalence spo₂)
+ ; trans = ×-transitive
+ {_<₁_ = _<₁_} (isEquivalence spo₁)
+ (<-resp-≈ spo₁) (trans spo₁)
+ {_≤₂_ = _<₂_} (trans spo₂)
+ ; compare = ×-compare (Eq.sym spo₁) (compare spo₁)
+ (compare spo₂)
+ }
+ where open IsStrictTotalOrder
+
+------------------------------------------------------------------------
+-- "Packages" can also be combined.
+
+module _ {ℓ₁ ℓ₂ ℓ₃ ℓ₄} where
+
+ ×-preorder : Preorder ℓ₁ ℓ₂ _ → Preorder ℓ₃ ℓ₄ _ → Preorder _ _ _
+ ×-preorder p₁ p₂ = record
+ { isPreorder = ×-isPreorder (isPreorder p₁) (isPreorder p₂)
+ } where open Preorder
+
+ ×-strictPartialOrder :
+ StrictPartialOrder ℓ₁ ℓ₂ _ → StrictPartialOrder ℓ₃ ℓ₄ _ →
+ StrictPartialOrder _ _ _
+ ×-strictPartialOrder s₁ s₂ = record
+ { isStrictPartialOrder = ×-isStrictPartialOrder
+ (isStrictPartialOrder s₁) (isStrictPartialOrder s₂)
+ } where open StrictPartialOrder
+
+ ×-strictTotalOrder :
+ StrictTotalOrder ℓ₁ ℓ₂ _ → StrictTotalOrder ℓ₃ ℓ₄ _ →
+ StrictTotalOrder _ _ _
+ ×-strictTotalOrder s₁ s₂ = record
+ { isStrictTotalOrder = ×-isStrictTotalOrder
+ (isStrictTotalOrder s₁) (isStrictTotalOrder s₂)
+ } where open StrictTotalOrder
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+_×-irreflexive_ = ×-irreflexive
+{-# WARNING_ON_USAGE _×-irreflexive_
+"Warning: _×-irreflexive_ was deprecated in v0.15.
+Please use ×-irreflexive instead."
+#-}
+_×-isPreorder_ = ×-isPreorder
+{-# WARNING_ON_USAGE _×-isPreorder_
+"Warning: _×-isPreorder_ was deprecated in v0.15.
+Please use ×-isPreorder instead."
+#-}
+_×-isStrictPartialOrder_ = ×-isStrictPartialOrder
+{-# WARNING_ON_USAGE _×-isStrictPartialOrder_
+"Warning: _×-isStrictPartialOrder_ was deprecated in v0.15.
+Please use ×-isStrictPartialOrder instead."
+#-}
+_×-isStrictTotalOrder_ = ×-isStrictTotalOrder
+{-# WARNING_ON_USAGE _×-isStrictTotalOrder_
+"Warning: _×-isStrictTotalOrder_ was deprecated in v0.15.
+Please use ×-isStrictTotalOrder instead."
+#-}
+_×-preorder_ = ×-preorder
+{-# WARNING_ON_USAGE _×-preorder_
+"Warning: _×-preorder_ was deprecated in v0.15.
+Please use ×-preorder instead."
+#-}
+_×-strictPartialOrder_ = ×-strictPartialOrder
+{-# WARNING_ON_USAGE _×-strictPartialOrder_
+"Warning: _×-strictPartialOrder_ was deprecated in v0.15.
+Please use ×-strictPartialOrder instead."
+#-}
+_×-strictTotalOrder_ = ×-strictTotalOrder
+{-# WARNING_ON_USAGE _×-strictTotalOrder_
+"Warning: _×-strictTotalOrder_ was deprecated in v0.15.
+Please use ×-strictTotalOrder instead."
+#-}
+×-≈-respects₂ = ×-respects₂
+{-# WARNING_ON_USAGE ×-≈-respects₂
+"Warning: ×-≈-respects₂ was deprecated in v0.15.
+Please use ×-respects₂ instead."
+#-}
diff --git a/src/Data/Product/Relation/Pointwise/Dependent.agda b/src/Data/Product/Relation/Pointwise/Dependent.agda
new file mode 100644
index 0000000..c3fe67f
--- /dev/null
+++ b/src/Data/Product/Relation/Pointwise/Dependent.agda
@@ -0,0 +1,449 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Pointwise lifting of binary relations to sigma types
+------------------------------------------------------------------------
+
+module Data.Product.Relation.Pointwise.Dependent where
+
+open import Data.Product as Prod
+open import Level
+open import Function
+open import Function.Equality as F using (_⟶_; _⟨$⟩_)
+open import Function.Equivalence as Eq
+ using (Equivalence; _⇔_; module Equivalence)
+open import Function.Injection as Inj
+ using (Injection; _↣_; module Injection; Injective)
+open import Function.Inverse as Inv
+ using (Inverse; _↔_; module Inverse)
+open import Function.LeftInverse as LeftInv
+ using (LeftInverse; _↞_; module LeftInverse;
+ _LeftInverseOf_; _RightInverseOf_)
+open import Function.Related as Related
+ using (_∼[_]_; lam; app-←; app-↢)
+open import Function.Surjection as Surj
+ using (Surjection; _↠_; module Surjection)
+open import Relation.Binary as B
+ using (_⇒_; Setoid; IsEquivalence)
+open import Relation.Binary.Indexed.Heterogeneous as I
+ using (IREL; IRel; IndexedSetoid; IsIndexedEquivalence)
+open import Relation.Binary.Indexed.Heterogeneous.Construct.At
+ using (_atₛ_)
+open import Relation.Binary.HeterogeneousEquality as H using (_≅_)
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+
+------------------------------------------------------------------------
+-- Pointwise lifting
+
+infixr 4 _,_
+
+record REL {a₁ a₂ b₁ b₂ ℓ₁ ℓ₂}
+ {A₁ : Set a₁} (B₁ : A₁ → Set b₁)
+ {A₂ : Set a₂} (B₂ : A₂ → Set b₂)
+ (_R₁_ : B.REL A₁ A₂ ℓ₁) (_R₂_ : IREL B₁ B₂ ℓ₂)
+ (xy₁ : Σ A₁ B₁) (xy₂ : Σ A₂ B₂)
+ : Set (a₁ ⊔ a₂ ⊔ b₁ ⊔ b₂ ⊔ ℓ₁ ⊔ ℓ₂) where
+ constructor _,_
+ field
+ proj₁ : (proj₁ xy₁) R₁ (proj₁ xy₂)
+ proj₂ : (proj₂ xy₁) R₂ (proj₂ xy₂)
+
+open REL public
+
+Pointwise : ∀ {a b ℓ₁ ℓ₂} {A : Set a} (B : A → Set b)
+ (_R₁_ : B.Rel A ℓ₁) (_R₂_ : IRel B ℓ₂) → B.Rel (Σ A B) _
+Pointwise B = REL B B
+
+------------------------------------------------------------------------
+-- Pointwise preserves many relational properties
+
+module _ {a b ℓ₁ ℓ₂} {A : Set a} {B : A → Set b}
+ {R₁ : B.Rel A ℓ₁} {R₂ : IRel B ℓ₂} where
+
+ refl : B.Reflexive R₁ → I.Reflexive B R₂ →
+ B.Reflexive (Pointwise B R₁ R₂)
+ refl refl₁ refl₂ = (refl₁ , refl₂)
+
+ symmetric : B.Symmetric R₁ → I.Symmetric B R₂ →
+ B.Symmetric (Pointwise B R₁ R₂)
+ symmetric sym₁ sym₂ (x₁Rx₂ , y₁Ry₂) = (sym₁ x₁Rx₂ , sym₂ y₁Ry₂)
+
+ transitive : B.Transitive R₁ → I.Transitive B R₂ →
+ B.Transitive (Pointwise B R₁ R₂)
+ transitive trans₁ trans₂ (x₁Rx₂ , y₁Ry₂) (x₂Rx₃ , y₂Ry₃) =
+ (trans₁ x₁Rx₂ x₂Rx₃ , trans₂ y₁Ry₂ y₂Ry₃)
+
+ isEquivalence : IsEquivalence R₁ → IsIndexedEquivalence B R₂ →
+ IsEquivalence (Pointwise B R₁ R₂)
+ isEquivalence eq₁ eq₂ = record
+ { refl = refl (IsEquivalence.refl eq₁)
+ (IsIndexedEquivalence.refl eq₂)
+ ; sym = symmetric (IsEquivalence.sym eq₁)
+ (IsIndexedEquivalence.sym eq₂)
+ ; trans = transitive (IsEquivalence.trans eq₁)
+ (IsIndexedEquivalence.trans eq₂)
+ }
+
+setoid : ∀ {b₁ b₂ i₁ i₂} → (A : Setoid b₁ b₂) →
+ IndexedSetoid (Setoid.Carrier A) i₁ i₂ →
+ B.Setoid _ _
+setoid s₁ s₂ = record
+ { isEquivalence = isEquivalence (Setoid.isEquivalence s₁)
+ (IndexedSetoid.isEquivalence s₂)
+ }
+
+------------------------------------------------------------------------
+-- The propositional equality setoid over sigma types can be
+-- decomposed using Pointwise
+
+module _ {a b} {A : Set a} {B : A → Set b} where
+
+ Pointwise-≡⇒≡ : Pointwise B _≡_ (λ x y → x ≅ y) ⇒ _≡_
+ Pointwise-≡⇒≡ (P.refl , H.refl) = P.refl
+
+ ≡⇒Pointwise-≡ : _≡_ ⇒ Pointwise B _≡_ (λ x y → x ≅ y)
+ ≡⇒Pointwise-≡ P.refl = (P.refl , H.refl)
+
+ Pointwise-≡↔≡ : Inverse (setoid (P.setoid A) (H.indexedSetoid B))
+ (P.setoid (Σ A B))
+ Pointwise-≡↔≡ = record
+ { to = record { _⟨$⟩_ = id; cong = Pointwise-≡⇒≡ }
+ ; from = record { _⟨$⟩_ = id; cong = ≡⇒Pointwise-≡ }
+ ; inverse-of = record
+ { left-inverse-of = uncurry (λ _ _ → (P.refl , H.refl))
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+------------------------------------------------------------------------
+-- Properties related to "relatedness"
+------------------------------------------------------------------------
+
+private
+
+ subst-cong : ∀ {i a p} {I : Set i} {A : I → Set a}
+ (P : ∀ {i} → A i → A i → Set p) {i i′} {x y : A i}
+ (i≡i′ : i P.≡ i′) →
+ P x y → P (P.subst A i≡i′ x) (P.subst A i≡i′ y)
+ subst-cong P P.refl p = p
+
+⟶ : ∀ {a₁ a₂ b₁ b₁′ b₂ b₂′}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : IndexedSetoid A₁ b₁ b₁′} (B₂ : IndexedSetoid A₂ b₂ b₂′)
+ (f : A₁ → A₂) → (∀ {x} → (B₁ atₛ x) ⟶ (B₂ atₛ (f x))) →
+ setoid (P.setoid A₁) B₁ ⟶ setoid (P.setoid A₂) B₂
+⟶ {A₁ = A₁} {A₂} {B₁} B₂ f g = record
+ { _⟨$⟩_ = fg
+ ; cong = fg-cong
+ }
+ where
+ open B.Setoid (setoid (P.setoid A₁) B₁)
+ using () renaming (_≈_ to _≈₁_)
+ open B.Setoid (setoid (P.setoid A₂) B₂)
+ using () renaming (_≈_ to _≈₂_)
+ open B using (_=[_]⇒_)
+
+ fg = Prod.map f (_⟨$⟩_ g)
+
+ fg-cong : _≈₁_ =[ fg ]⇒ _≈₂_
+ fg-cong (P.refl , ∼) = (P.refl , F.cong g ∼)
+
+
+module _ {a₁ a₂ b₁ b₁′ b₂ b₂′} {A₁ : Set a₁} {A₂ : Set a₂} where
+
+ equivalence : {B₁ : IndexedSetoid A₁ b₁ b₁′} {B₂ : IndexedSetoid A₂ b₂ b₂′}
+ (A₁⇔A₂ : A₁ ⇔ A₂) →
+ (∀ {x} → _⟶_ (B₁ atₛ x) (B₂ atₛ (Equivalence.to A₁⇔A₂ ⟨$⟩ x))) →
+ (∀ {y} → _⟶_ (B₂ atₛ y) (B₁ atₛ (Equivalence.from A₁⇔A₂ ⟨$⟩ y))) →
+ Equivalence (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+ equivalence {B₁} {B₂} A₁⇔A₂ B-to B-from = record
+ { to = ⟶ B₂ (_⟨$⟩_ (to A₁⇔A₂)) B-to
+ ; from = ⟶ B₁ (_⟨$⟩_ (from A₁⇔A₂)) B-from
+ } where open Equivalence
+
+ equivalence-↞ : (B₁ : IndexedSetoid A₁ b₁ b₁′) {B₂ : IndexedSetoid A₂ b₂ b₂′}
+ (A₁↞A₂ : A₁ ↞ A₂) →
+ (∀ {x} → Equivalence (B₁ atₛ (LeftInverse.from A₁↞A₂ ⟨$⟩ x))
+ (B₂ atₛ x)) →
+ Equivalence (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+ equivalence-↞ B₁ {B₂} A₁↞A₂ B₁⇔B₂ =
+ equivalence (LeftInverse.equivalence A₁↞A₂) B-to B-from
+ where
+ B-to : ∀ {x} → _⟶_ (B₁ atₛ x) (B₂ atₛ (LeftInverse.to A₁↞A₂ ⟨$⟩ x))
+ B-to = record
+ { _⟨$⟩_ = λ x → Equivalence.to B₁⇔B₂ ⟨$⟩
+ P.subst (IndexedSetoid.Carrier B₁)
+ (P.sym $ LeftInverse.left-inverse-of A₁↞A₂ _)
+ x
+ ; cong = F.cong (Equivalence.to B₁⇔B₂) ∘
+ subst-cong (λ {x} → IndexedSetoid._≈_ B₁ {x} {x})
+ (P.sym (LeftInverse.left-inverse-of A₁↞A₂ _))
+ }
+
+ B-from : ∀ {y} → _⟶_ (B₂ atₛ y) (B₁ atₛ (LeftInverse.from A₁↞A₂ ⟨$⟩ y))
+ B-from = Equivalence.from B₁⇔B₂
+
+ equivalence-↠ : {B₁ : IndexedSetoid A₁ b₁ b₁′} (B₂ : IndexedSetoid A₂ b₂ b₂′)
+ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → Equivalence (B₁ atₛ x) (B₂ atₛ (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
+ Equivalence (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+ equivalence-↠ {B₁ = B₁} B₂ A₁↠A₂ B₁⇔B₂ =
+ equivalence (Surjection.equivalence A₁↠A₂) B-to B-from
+ where
+ B-to : ∀ {x} → _⟶_ (B₁ atₛ x) (B₂ atₛ (Surjection.to A₁↠A₂ ⟨$⟩ x))
+ B-to = Equivalence.to B₁⇔B₂
+
+ B-from : ∀ {y} → _⟶_ (B₂ atₛ y) (B₁ atₛ (Surjection.from A₁↠A₂ ⟨$⟩ y))
+ B-from = record
+ { _⟨$⟩_ = λ x → Equivalence.from B₁⇔B₂ ⟨$⟩
+ P.subst (IndexedSetoid.Carrier B₂)
+ (P.sym $ Surjection.right-inverse-of A₁↠A₂ _)
+ x
+ ; cong = F.cong (Equivalence.from B₁⇔B₂) ∘
+ subst-cong (λ {x} → IndexedSetoid._≈_ B₂ {x} {x})
+ (P.sym (Surjection.right-inverse-of A₁↠A₂ _))
+ }
+
+ injection : {B₁ : IndexedSetoid A₁ b₁ b₁′} (B₂ : IndexedSetoid A₂ b₂ b₂′) →
+ (A₁↣A₂ : A₁ ↣ A₂) →
+ (∀ {x} → Injection (B₁ atₛ x) (B₂ atₛ (Injection.to A₁↣A₂ ⟨$⟩ x))) →
+ Injection (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+ injection {B₁ = B₁} B₂ A₁↣A₂ B₁↣B₂ = record
+ { to = to
+ ; injective = inj
+ }
+ where
+ to = ⟶ B₂ (Injection.to A₁↣A₂ ⟨$⟩_) (Injection.to B₁↣B₂)
+
+ inj : Injective to
+ inj (x , y) =
+ Injection.injective A₁↣A₂ x ,
+ lemma (Injection.injective A₁↣A₂ x) y
+ where
+ lemma :
+ ∀ {x x′}
+ {y : IndexedSetoid.Carrier B₁ x} {y′ : IndexedSetoid.Carrier B₁ x′} →
+ x ≡ x′ →
+ (eq : IndexedSetoid._≈_ B₂ (Injection.to B₁↣B₂ ⟨$⟩ y)
+ (Injection.to B₁↣B₂ ⟨$⟩ y′)) →
+ IndexedSetoid._≈_ B₁ y y′
+ lemma P.refl = Injection.injective B₁↣B₂
+
+ left-inverse : (B₁ : IndexedSetoid A₁ b₁ b₁′) {B₂ : IndexedSetoid A₂ b₂ b₂′} →
+ (A₁↞A₂ : A₁ ↞ A₂) →
+ (∀ {x} → LeftInverse (B₁ atₛ (LeftInverse.from A₁↞A₂ ⟨$⟩ x))
+ (B₂ atₛ x)) →
+ LeftInverse (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+ left-inverse B₁ {B₂} A₁↞A₂ B₁↞B₂ = record
+ { to = Equivalence.to eq
+ ; from = Equivalence.from eq
+ ; left-inverse-of = left
+ }
+ where
+ eq = equivalence-↞ B₁ A₁↞A₂ (LeftInverse.equivalence B₁↞B₂)
+
+ left : Equivalence.from eq LeftInverseOf Equivalence.to eq
+ left (x , y) =
+ LeftInverse.left-inverse-of A₁↞A₂ x ,
+ IndexedSetoid.trans B₁
+ (LeftInverse.left-inverse-of B₁↞B₂ _)
+ (lemma (P.sym (LeftInverse.left-inverse-of A₁↞A₂ x)))
+ where
+ lemma :
+ ∀ {x x′ y} (eq : x ≡ x′) →
+ IndexedSetoid._≈_ B₁ (P.subst (IndexedSetoid.Carrier B₁) eq y) y
+ lemma P.refl = IndexedSetoid.refl B₁
+
+ surjection : {B₁ : IndexedSetoid A₁ b₁ b₁′} (B₂ : IndexedSetoid A₂ b₂ b₂′) →
+ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → Surjection (B₁ atₛ x) (B₂ atₛ (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
+ Surjection (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+ surjection B₂ A₁↠A₂ B₁↠B₂ = record
+ { to = Equivalence.to eq
+ ; surjective = record
+ { from = Equivalence.from eq
+ ; right-inverse-of = right
+ }
+ }
+ where
+ eq = equivalence-↠ B₂ A₁↠A₂ (Surjection.equivalence B₁↠B₂)
+
+ right : Equivalence.from eq RightInverseOf Equivalence.to eq
+ right (x , y) =
+ Surjection.right-inverse-of A₁↠A₂ x ,
+ IndexedSetoid.trans B₂
+ (Surjection.right-inverse-of B₁↠B₂ _)
+ (lemma (P.sym $ Surjection.right-inverse-of A₁↠A₂ x))
+ where
+ lemma : ∀ {x x′ y} (eq : x ≡ x′) →
+ IndexedSetoid._≈_ B₂ (P.subst (IndexedSetoid.Carrier B₂) eq y) y
+ lemma P.refl = IndexedSetoid.refl B₂
+
+ inverse : {B₁ : IndexedSetoid A₁ b₁ b₁′} (B₂ : IndexedSetoid A₂ b₂ b₂′) →
+ (A₁↔A₂ : A₁ ↔ A₂) →
+ (∀ {x} → Inverse (B₁ atₛ x) (B₂ atₛ (Inverse.to A₁↔A₂ ⟨$⟩ x))) →
+ Inverse (setoid (P.setoid A₁) B₁) (setoid (P.setoid A₂) B₂)
+ inverse {B₁} B₂ A₁↔A₂ B₁↔B₂ = record
+ { to = Surjection.to surj
+ ; from = Surjection.from surj
+ ; inverse-of = record
+ { left-inverse-of = left
+ ; right-inverse-of = Surjection.right-inverse-of surj
+ }
+ }
+ where
+ surj = surjection B₂ (Inverse.surjection A₁↔A₂)
+ (Inverse.surjection B₁↔B₂)
+
+ left : Surjection.from surj LeftInverseOf Surjection.to surj
+ left (x , y) =
+ Inverse.left-inverse-of A₁↔A₂ x ,
+ IndexedSetoid.trans B₁
+ (lemma (P.sym (Inverse.left-inverse-of A₁↔A₂ x))
+ (P.sym (Inverse.right-inverse-of A₁↔A₂
+ (Inverse.to A₁↔A₂ ⟨$⟩ x))))
+ (Inverse.left-inverse-of B₁↔B₂ y)
+ where
+ lemma :
+ ∀ {x x′ y} → x ≡ x′ →
+ (eq : (Inverse.to A₁↔A₂ ⟨$⟩ x) ≡ (Inverse.to A₁↔A₂ ⟨$⟩ x′)) →
+ IndexedSetoid._≈_ B₁
+ (Inverse.from B₁↔B₂ ⟨$⟩ P.subst (IndexedSetoid.Carrier B₂) eq y)
+ (Inverse.from B₁↔B₂ ⟨$⟩ y)
+ lemma P.refl P.refl = IndexedSetoid.refl B₁
+
+------------------------------------------------------------------------
+
+module _ {a₁ a₂} {A₁ : Set a₁} {A₂ : Set a₂}
+ {b₁ b₂} {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ where
+
+ ⇔ : (A₁⇔A₂ : A₁ ⇔ A₂) →
+ (∀ {x} → B₁ x → B₂ (Equivalence.to A₁⇔A₂ ⟨$⟩ x)) →
+ (∀ {y} → B₂ y → B₁ (Equivalence.from A₁⇔A₂ ⟨$⟩ y)) →
+ Σ A₁ B₁ ⇔ Σ A₂ B₂
+ ⇔ A₁⇔A₂ B-to B-from =
+ Inverse.equivalence Pointwise-≡↔≡ ⟨∘⟩
+ equivalence A₁⇔A₂
+ (Inverse.to (H.≡↔≅ B₂) ⊚ P.→-to-⟶ B-to ⊚ Inverse.from (H.≡↔≅ B₁))
+ (Inverse.to (H.≡↔≅ B₁) ⊚ P.→-to-⟶ B-from ⊚ Inverse.from (H.≡↔≅ B₂))
+ ⟨∘⟩
+ Eq.sym (Inverse.equivalence (Pointwise-≡↔≡ {B = B₁}))
+ where
+ open Eq using () renaming (_∘_ to _⟨∘⟩_)
+ open F using () renaming (_∘_ to _⊚_)
+
+ ⇔-↠ : ∀ (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → _⇔_ (B₁ x) (B₂ (Surjection.to A₁↠A₂ ⟨$⟩ x))) →
+ _⇔_ (Σ A₁ B₁) (Σ A₂ B₂)
+ ⇔-↠ A₁↠A₂ B₁⇔B₂ =
+ Inverse.equivalence Pointwise-≡↔≡ ⟨∘⟩
+ equivalence-↠ (H.indexedSetoid B₂) A₁↠A₂
+ (Inverse.equivalence (H.≡↔≅ B₂) ⟨∘⟩
+ B₁⇔B₂ ⟨∘⟩
+ Inverse.equivalence (Inv.sym (H.≡↔≅ B₁))) ⟨∘⟩
+ Eq.sym (Inverse.equivalence Pointwise-≡↔≡)
+ where open Eq using () renaming (_∘_ to _⟨∘⟩_)
+
+ ↣ : ∀ (A₁↣A₂ : A₁ ↣ A₂) →
+ (∀ {x} → B₁ x ↣ B₂ (Injection.to A₁↣A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ↣ Σ A₂ B₂
+ ↣ A₁↣A₂ B₁↣B₂ =
+ Inverse.injection Pointwise-≡↔≡ ⟨∘⟩
+ injection (H.indexedSetoid B₂) A₁↣A₂
+ (Inverse.injection (H.≡↔≅ B₂) ⟨∘⟩
+ B₁↣B₂ ⟨∘⟩
+ Inverse.injection (Inv.sym (H.≡↔≅ B₁))) ⟨∘⟩
+ Inverse.injection (Inv.sym Pointwise-≡↔≡)
+ where open Inj using () renaming (_∘_ to _⟨∘⟩_)
+
+ ↞ : (A₁↞A₂ : A₁ ↞ A₂) →
+ (∀ {x} → B₁ (LeftInverse.from A₁↞A₂ ⟨$⟩ x) ↞ B₂ x) →
+ Σ A₁ B₁ ↞ Σ A₂ B₂
+ ↞ A₁↞A₂ B₁↞B₂ =
+ Inverse.left-inverse Pointwise-≡↔≡ ⟨∘⟩
+ left-inverse (H.indexedSetoid B₁) A₁↞A₂
+ (Inverse.left-inverse (H.≡↔≅ B₂) ⟨∘⟩
+ B₁↞B₂ ⟨∘⟩
+ Inverse.left-inverse (Inv.sym (H.≡↔≅ B₁))) ⟨∘⟩
+ Inverse.left-inverse (Inv.sym Pointwise-≡↔≡)
+ where open LeftInv using () renaming (_∘_ to _⟨∘⟩_)
+
+ ↠ : (A₁↠A₂ : A₁ ↠ A₂) →
+ (∀ {x} → B₁ x ↠ B₂ (Surjection.to A₁↠A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ↠ Σ A₂ B₂
+ ↠ A₁↠A₂ B₁↠B₂ =
+ Inverse.surjection Pointwise-≡↔≡ ⟨∘⟩
+ surjection (H.indexedSetoid B₂) A₁↠A₂
+ (Inverse.surjection (H.≡↔≅ B₂) ⟨∘⟩
+ B₁↠B₂ ⟨∘⟩
+ Inverse.surjection (Inv.sym (H.≡↔≅ B₁))) ⟨∘⟩
+ Inverse.surjection (Inv.sym Pointwise-≡↔≡)
+ where open Surj using () renaming (_∘_ to _⟨∘⟩_)
+
+ ↔ : (A₁↔A₂ : A₁ ↔ A₂) →
+ (∀ {x} → B₁ x ↔ B₂ (Inverse.to A₁↔A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ↔ Σ A₂ B₂
+ ↔ A₁↔A₂ B₁↔B₂ =
+ Pointwise-≡↔≡ ⟨∘⟩
+ inverse (H.indexedSetoid B₂) A₁↔A₂
+ (H.≡↔≅ B₂ ⟨∘⟩ B₁↔B₂ ⟨∘⟩ Inv.sym (H.≡↔≅ B₁)) ⟨∘⟩
+ Inv.sym Pointwise-≡↔≡
+ where open Inv using () renaming (_∘_ to _⟨∘⟩_)
+
+private
+
+ swap-coercions : ∀ {k a₁ a₂ b₁ b₂} {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} (B₂ : A₂ → Set b₂)
+ (A₁↔A₂ : _↔_ A₁ A₂) →
+ (∀ {x} → B₁ x ∼[ k ] B₂ (Inverse.to A₁↔A₂ ⟨$⟩ x)) →
+ ∀ {x} → B₁ (Inverse.from A₁↔A₂ ⟨$⟩ x) ∼[ k ] B₂ x
+ swap-coercions {k} {B₁ = B₁} B₂ A₁↔A₂ eq {x} =
+ B₁ (Inverse.from A₁↔A₂ ⟨$⟩ x)
+ ∼⟨ eq ⟩
+ B₂ (Inverse.to A₁↔A₂ ⟨$⟩ (Inverse.from A₁↔A₂ ⟨$⟩ x))
+ ↔⟨ Related.K-reflexive
+ (P.cong B₂ $ Inverse.right-inverse-of A₁↔A₂ x) ⟩
+ B₂ x
+ ∎
+ where open Related.EquationalReasoning
+
+cong : ∀ {k a₁ a₂ b₁ b₂}
+ {A₁ : Set a₁} {A₂ : Set a₂}
+ {B₁ : A₁ → Set b₁} {B₂ : A₂ → Set b₂}
+ (A₁↔A₂ : _↔_ A₁ A₂) →
+ (∀ {x} → B₁ x ∼[ k ] B₂ (Inverse.to A₁↔A₂ ⟨$⟩ x)) →
+ Σ A₁ B₁ ∼[ k ] Σ A₂ B₂
+cong {Related.implication} =
+ λ A₁↔A₂ → Prod.map (_⟨$⟩_ (Inverse.to A₁↔A₂))
+cong {Related.reverse-implication} {B₂ = B₂} =
+ λ A₁↔A₂ B₁←B₂ → lam (Prod.map (_⟨$⟩_ (Inverse.from A₁↔A₂))
+ (app-← (swap-coercions B₂ A₁↔A₂ B₁←B₂)))
+cong {Related.equivalence} = ⇔-↠ ∘ Inverse.surjection
+cong {Related.injection} = ↣ ∘ Inverse.injection
+cong {Related.reverse-injection} {B₂ = B₂} =
+ λ A₁↔A₂ B₁↢B₂ → lam (↣ (Inverse.injection (Inv.sym A₁↔A₂))
+ (app-↢ (swap-coercions B₂ A₁↔A₂ B₁↢B₂)))
+cong {Related.left-inverse} =
+ λ A₁↔A₂ → ↞ (Inverse.left-inverse A₁↔A₂) ∘ swap-coercions _ A₁↔A₂
+cong {Related.surjection} = ↠ ∘ Inverse.surjection
+cong {Related.bijection} = ↔
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+Rel = Pointwise
+{-# WARNING_ON_USAGE Rel
+"Warning: Rel was deprecated in v0.15.
+Please use Pointwise instead."
+#-}
+Rel↔≡ = Pointwise-≡↔≡
+{-# WARNING_ON_USAGE Rel↔≡
+"Warning: Rel↔≡ was deprecated in v0.15.
+Please use Pointwise-≡↔≡ instead."
+#-}
diff --git a/src/Data/Product/Relation/Pointwise/NonDependent.agda b/src/Data/Product/Relation/Pointwise/NonDependent.agda
new file mode 100644
index 0000000..bff2a2b
--- /dev/null
+++ b/src/Data/Product/Relation/Pointwise/NonDependent.agda
@@ -0,0 +1,499 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Pointwise products of binary relations
+------------------------------------------------------------------------
+
+module Data.Product.Relation.Pointwise.NonDependent where
+
+open import Data.Product as Prod
+import Data.Product.Relation.Pointwise.Dependent as Dependent
+open import Data.Sum
+open import Data.Unit.Base using (⊤)
+open import Function
+open import Function.Equality as F using (_⟶_; _⟨$⟩_)
+open import Function.Equivalence as Eq
+ using (Equivalence; _⇔_; module Equivalence)
+open import Function.Injection as Inj
+ using (Injection; _↣_; module Injection)
+open import Function.Inverse as Inv
+ using (Inverse; _↔_; module Inverse)
+open import Function.LeftInverse as LeftInv
+ using (LeftInverse; _↞_; _LeftInverseOf_; module LeftInverse)
+open import Function.Related
+open import Function.Surjection as Surj
+ using (Surjection; _↠_; module Surjection)
+import Relation.Nullary.Decidable as Dec
+open import Relation.Nullary.Product
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+
+module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where
+
+------------------------------------------------------------------------
+-- Pointwise lifting
+
+ Pointwise : Rel A₁ ℓ₁ → Rel A₂ ℓ₂ → Rel (A₁ × A₂) _
+ Pointwise _∼₁_ _∼₂_ = (_∼₁_ on proj₁) -×- (_∼₂_ on proj₂)
+
+------------------------------------------------------------------------
+-- Pointwise preserves many relational properties
+
+ ×-reflexive : ∀ {_≈₁_ _∼₁_ _≈₂_ _∼₂_} →
+ _≈₁_ ⇒ _∼₁_ → _≈₂_ ⇒ _∼₂_ →
+ (Pointwise _≈₁_ _≈₂_) ⇒ (Pointwise _∼₁_ _∼₂_)
+ ×-reflexive refl₁ refl₂ (x∼y₁ , x∼y₂) = refl₁ x∼y₁ , refl₂ x∼y₂
+
+ ×-refl : ∀ {_∼₁_ _∼₂_} →
+ Reflexive _∼₁_ → Reflexive _∼₂_ →
+ Reflexive (Pointwise _∼₁_ _∼₂_)
+ ×-refl refl₁ refl₂ = refl₁ , refl₂
+
+ ×-irreflexive₁ : ∀ {_≈₁_ _<₁_ _≈₂_ _<₂_} →
+ Irreflexive _≈₁_ _<₁_ →
+ Irreflexive (Pointwise _≈₁_ _≈₂_) (Pointwise _<₁_ _<₂_)
+ ×-irreflexive₁ ir x≈y x<y = ir (proj₁ x≈y) (proj₁ x<y)
+
+ ×-irreflexive₂ : ∀ {_≈₁_ _<₁_ _≈₂_ _<₂_} →
+ Irreflexive _≈₂_ _<₂_ →
+ Irreflexive (Pointwise _≈₁_ _≈₂_) (Pointwise _<₁_ _<₂_)
+ ×-irreflexive₂ ir x≈y x<y = ir (proj₂ x≈y) (proj₂ x<y)
+
+ ×-symmetric : ∀ {_∼₁_ _∼₂_} → Symmetric _∼₁_ → Symmetric _∼₂_ →
+ Symmetric (Pointwise _∼₁_ _∼₂_)
+ ×-symmetric sym₁ sym₂ (x∼y₁ , x∼y₂) = sym₁ x∼y₁ , sym₂ x∼y₂
+
+ ×-transitive : ∀ {_∼₁_ _∼₂_} → Transitive _∼₁_ → Transitive _∼₂_ →
+ Transitive (Pointwise _∼₁_ _∼₂_)
+ ×-transitive trans₁ trans₂ x∼y y∼z =
+ trans₁ (proj₁ x∼y) (proj₁ y∼z) ,
+ trans₂ (proj₂ x∼y) (proj₂ y∼z)
+
+ ×-antisymmetric : ∀ {_≈₁_ _≤₁_ _≈₂_ _≤₂_} →
+ Antisymmetric _≈₁_ _≤₁_ → Antisymmetric _≈₂_ _≤₂_ →
+ Antisymmetric (Pointwise _≈₁_ _≈₂_) (Pointwise _≤₁_ _≤₂_)
+ ×-antisymmetric antisym₁ antisym₂ (x≤y₁ , x≤y₂) (y≤x₁ , y≤x₂) =
+ (antisym₁ x≤y₁ y≤x₁ , antisym₂ x≤y₂ y≤x₂)
+
+ ×-asymmetric₁ : ∀ {_<₁_ _∼₂_} → Asymmetric _<₁_ →
+ Asymmetric (Pointwise _<₁_ _∼₂_)
+ ×-asymmetric₁ asym₁ x<y y<x = asym₁ (proj₁ x<y) (proj₁ y<x)
+
+ ×-asymmetric₂ : ∀ {_∼₁_ _<₂_} → Asymmetric _<₂_ →
+ Asymmetric (Pointwise _∼₁_ _<₂_)
+ ×-asymmetric₂ asym₂ x<y y<x = asym₂ (proj₂ x<y) (proj₂ y<x)
+
+ ×-respects₂ : ∀ {_≈₁_ _∼₁_ _≈₂_ _∼₂_} →
+ _∼₁_ Respects₂ _≈₁_ → _∼₂_ Respects₂ _≈₂_ →
+ (Pointwise _∼₁_ _∼₂_) Respects₂ (Pointwise _≈₁_ _≈₂_)
+ ×-respects₂ {_≈₁_} {_∼₁_} {_≈₂_} {_∼₂_} resp₁ resp₂ = resp¹ , resp²
+ where
+ _∼_ = Pointwise _∼₁_ _∼₂_
+ _≈_ = Pointwise _≈₁_ _≈₂_
+
+ resp¹ : ∀ {x} → (x ∼_) Respects _≈_
+ resp¹ y≈y' x∼y = proj₁ resp₁ (proj₁ y≈y') (proj₁ x∼y) ,
+ proj₁ resp₂ (proj₂ y≈y') (proj₂ x∼y)
+
+ resp² : ∀ {y} → (_∼ y) Respects _≈_
+ resp² x≈x' x∼y = proj₂ resp₁ (proj₁ x≈x') (proj₁ x∼y) ,
+ proj₂ resp₂ (proj₂ x≈x') (proj₂ x∼y)
+
+ ×-total : ∀ {_∼₁_ _∼₂_} → Symmetric _∼₁_ →
+ Total _∼₁_ → Total _∼₂_ →
+ Total (Pointwise _∼₁_ _∼₂_)
+ ×-total sym₁ total₁ total₂ (x₁ , x₂) (y₁ , y₂)
+ with total₁ x₁ y₁ | total₂ x₂ y₂
+ ... | inj₁ x₁∼y₁ | inj₁ x₂∼y₂ = inj₁ ( x₁∼y₁ , x₂∼y₂)
+ ... | inj₁ x₁∼y₁ | inj₂ y₂∼x₂ = inj₂ (sym₁ x₁∼y₁ , y₂∼x₂)
+ ... | inj₂ y₁∼x₁ | inj₂ y₂∼x₂ = inj₂ ( y₁∼x₁ , y₂∼x₂)
+ ... | inj₂ y₁∼x₁ | inj₁ x₂∼y₂ = inj₁ (sym₁ y₁∼x₁ , x₂∼y₂)
+
+ ×-decidable : ∀ {_∼₁_ _∼₂_} →
+ Decidable _∼₁_ → Decidable _∼₂_ →
+ Decidable (Pointwise _∼₁_ _∼₂_)
+ ×-decidable _≟₁_ _≟₂_ (x₁ , x₂) (y₁ , y₂) =
+ (x₁ ≟₁ y₁) ×-dec (x₂ ≟₂ y₂)
+
+ -- Some collections of properties which are preserved by ×-Rel.
+
+ ×-isEquivalence : ∀ {_≈₁_ _≈₂_} →
+ IsEquivalence _≈₁_ → IsEquivalence _≈₂_ →
+ IsEquivalence (Pointwise _≈₁_ _≈₂_)
+ ×-isEquivalence {_≈₁_ = _≈₁_} {_≈₂_ = _≈₂_} eq₁ eq₂ = record
+ { refl = ×-refl {_∼₁_ = _≈₁_} {_∼₂_ = _≈₂_}
+ (refl eq₁) (refl eq₂)
+ ; sym = ×-symmetric {_∼₁_ = _≈₁_} {_∼₂_ = _≈₂_}
+ (sym eq₁) (sym eq₂)
+ ; trans = ×-transitive {_∼₁_ = _≈₁_} {_∼₂_ = _≈₂_}
+ (trans eq₁) (trans eq₂)
+ }
+ where open IsEquivalence
+
+ ×-isDecEquivalence : ∀ {_≈₁_ _≈₂_} →
+ IsDecEquivalence _≈₁_ → IsDecEquivalence _≈₂_ →
+ IsDecEquivalence (Pointwise _≈₁_ _≈₂_)
+ ×-isDecEquivalence eq₁ eq₂ = record
+ { isEquivalence = ×-isEquivalence
+ (isEquivalence eq₁) (isEquivalence eq₂)
+ ; _≟_ = ×-decidable (_≟_ eq₁) (_≟_ eq₂)
+ }
+ where open IsDecEquivalence
+
+ ×-isPreorder : ∀ {_≈₁_ _∼₁_ _≈₂_ _∼₂_} →
+ IsPreorder _≈₁_ _∼₁_ → IsPreorder _≈₂_ _∼₂_ →
+ IsPreorder (Pointwise _≈₁_ _≈₂_) (Pointwise _∼₁_ _∼₂_)
+ ×-isPreorder {_∼₁_ = _∼₁_} {_∼₂_ = _∼₂_} pre₁ pre₂ = record
+ { isEquivalence = ×-isEquivalence
+ (isEquivalence pre₁) (isEquivalence pre₂)
+ ; reflexive = ×-reflexive {_∼₁_ = _∼₁_} {_∼₂_ = _∼₂_}
+ (reflexive pre₁) (reflexive pre₂)
+ ; trans = ×-transitive {_∼₁_ = _∼₁_} {_∼₂_ = _∼₂_}
+ (trans pre₁) (trans pre₂)
+ }
+ where open IsPreorder
+
+ ×-isPartialOrder : ∀ {_≈₁_ _≤₁_ _≈₂_ _≤₂_} →
+ IsPartialOrder _≈₁_ _≤₁_ → IsPartialOrder _≈₂_ _≤₂_ →
+ IsPartialOrder (Pointwise _≈₁_ _≈₂_) (Pointwise _≤₁_ _≤₂_)
+ ×-isPartialOrder {_≤₁_ = _≤₁_} {_≤₂_ = _≤₂_} po₁ po₂ = record
+ { isPreorder = ×-isPreorder (isPreorder po₁) (isPreorder po₂)
+ ; antisym = ×-antisymmetric {_≤₁_ = _≤₁_} {_≤₂_ = _≤₂_}
+ (antisym po₁) (antisym po₂)
+ }
+ where open IsPartialOrder
+
+ ×-isStrictPartialOrder : ∀ {_≈₁_ _<₁_ _≈₂_ _<₂_} →
+ IsStrictPartialOrder _≈₁_ _<₁_ → IsStrictPartialOrder _≈₂_ _<₂_ →
+ IsStrictPartialOrder (Pointwise _≈₁_ _≈₂_) (Pointwise _<₁_ _<₂_)
+ ×-isStrictPartialOrder {_<₁_ = _<₁_} {_≈₂_ = _≈₂_} {_<₂_ = _<₂_}
+ spo₁ spo₂ =
+ record
+ { isEquivalence = ×-isEquivalence
+ (isEquivalence spo₁) (isEquivalence spo₂)
+ ; irrefl = ×-irreflexive₁ {_<₁_ = _<₁_} {_≈₂_} {_<₂_}
+ (irrefl spo₁)
+ ; trans = ×-transitive {_∼₁_ = _<₁_} {_<₂_}
+ (trans spo₁) (trans spo₂)
+ ; <-resp-≈ = ×-respects₂ (<-resp-≈ spo₁) (<-resp-≈ spo₂)
+ }
+ where open IsStrictPartialOrder
+
+------------------------------------------------------------------------
+-- "Packages" can also be combined.
+
+module _ {ℓ₁ ℓ₂ ℓ₃ ℓ₄} where
+
+ ×-preorder : Preorder ℓ₁ ℓ₂ _ → Preorder ℓ₃ ℓ₄ _ → Preorder _ _ _
+ ×-preorder p₁ p₂ = record
+ { isPreorder = ×-isPreorder (isPreorder p₁) (isPreorder p₂)
+ } where open Preorder
+
+ ×-setoid : Setoid ℓ₁ ℓ₂ → Setoid ℓ₃ ℓ₄ → Setoid _ _
+ ×-setoid s₁ s₂ = record
+ { isEquivalence =
+ ×-isEquivalence (isEquivalence s₁) (isEquivalence s₂)
+ } where open Setoid
+
+ ×-decSetoid : DecSetoid ℓ₁ ℓ₂ → DecSetoid ℓ₃ ℓ₄ → DecSetoid _ _
+ ×-decSetoid s₁ s₂ = record
+ { isDecEquivalence =
+ ×-isDecEquivalence (isDecEquivalence s₁) (isDecEquivalence s₂)
+ } where open DecSetoid
+
+ ×-poset : Poset ℓ₁ ℓ₂ _ → Poset ℓ₃ ℓ₄ _ → Poset _ _ _
+ ×-poset s₁ s₂ = record
+ { isPartialOrder = ×-isPartialOrder (isPartialOrder s₁)
+ (isPartialOrder s₂)
+ } where open Poset
+
+ ×-strictPartialOrder :
+ StrictPartialOrder ℓ₁ ℓ₂ _ → StrictPartialOrder ℓ₃ ℓ₄ _ →
+ StrictPartialOrder _ _ _
+ ×-strictPartialOrder s₁ s₂ = record
+ { isStrictPartialOrder = ×-isStrictPartialOrder
+ (isStrictPartialOrder s₁)
+ (isStrictPartialOrder s₂)
+ } where open StrictPartialOrder
+
+ -- A piece of infix notation for combining setoids
+ infix 4 _×ₛ_
+ _×ₛ_ : Setoid ℓ₁ ℓ₂ → Setoid ℓ₃ ℓ₄ → Setoid _ _
+ _×ₛ_ = ×-setoid
+
+------------------------------------------------------------------------
+-- The propositional equality setoid over products can be
+-- decomposed using ×-Rel
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ ≡×≡⇒≡ : Pointwise _≡_ _≡_ ⇒ _≡_ {A = A × B}
+ ≡×≡⇒≡ (P.refl , P.refl) = P.refl
+
+ ≡⇒≡×≡ : _≡_ {A = A × B} ⇒ Pointwise _≡_ _≡_
+ ≡⇒≡×≡ P.refl = (P.refl , P.refl)
+
+ Pointwise-≡↔≡ : Inverse (×-setoid (P.setoid A) (P.setoid B))
+ (P.setoid (A × B))
+ Pointwise-≡↔≡ = record
+ { to = record { _⟨$⟩_ = id; cong = ≡×≡⇒≡ }
+ ; from = record { _⟨$⟩_ = id; cong = ≡⇒≡×≡ }
+ ; inverse-of = record
+ { left-inverse-of = λ _ → (P.refl , P.refl)
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+ ≡?×≡?⇒≡? : Decidable {A = A} _≡_ → Decidable {A = B} _≡_ →
+ Decidable {A = A × B} _≡_
+ ≡?×≡?⇒≡? ≟₁ ≟₂ p₁ p₂ =
+ Dec.map′ ≡×≡⇒≡ ≡⇒≡×≡ (×-decidable ≟₁ ≟₂ p₁ p₂)
+
+------------------------------------------------------------------------
+-- Some properties related to "relatedness"
+
+_×-⟶_ : ∀ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂} →
+ A ⟶ B → C ⟶ D → (A ×ₛ C) ⟶ (B ×ₛ D)
+_×-⟶_ {A = A} {B} {C} {D} f g = record
+ { _⟨$⟩_ = fg
+ ; cong = fg-cong
+ }
+ where
+ open Setoid (A ×ₛ C) using () renaming (_≈_ to _≈AC_)
+ open Setoid (B ×ₛ D) using () renaming (_≈_ to _≈BD_)
+
+ fg = Prod.map (f ⟨$⟩_) (g ⟨$⟩_)
+
+ fg-cong : _≈AC_ =[ fg ]⇒ _≈BD_
+ fg-cong (_∼₁_ , _∼₂_) = (F.cong f _∼₁_ , F.cong g _∼₂_)
+
+module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
+ where
+
+ _×-equivalence_ : Equivalence A B → Equivalence C D →
+ Equivalence (A ×ₛ C) (B ×ₛ D)
+ _×-equivalence_ A⇔B C⇔D = record
+ { to = to A⇔B ×-⟶ to C⇔D
+ ; from = from A⇔B ×-⟶ from C⇔D
+ } where open Equivalence
+
+ _×-injection_ : Injection A B → Injection C D →
+ Injection (A ×ₛ C) (B ×ₛ D)
+ A↣B ×-injection C↣D = record
+ { to = to A↣B ×-⟶ to C↣D
+ ; injective = Prod.map (injective A↣B) (injective C↣D)
+ } where open Injection
+
+ _×-left-inverse_ : LeftInverse A B → LeftInverse C D →
+ LeftInverse (A ×ₛ C) (B ×ₛ D)
+ A↞B ×-left-inverse C↞D = record
+ { to = Equivalence.to eq
+ ; from = Equivalence.from eq
+ ; left-inverse-of = left
+ }
+ where
+ open LeftInverse
+ eq = LeftInverse.equivalence A↞B ×-equivalence
+ LeftInverse.equivalence C↞D
+
+ left : Equivalence.from eq LeftInverseOf Equivalence.to eq
+ left (x , y) = (left-inverse-of A↞B x , left-inverse-of C↞D y)
+
+module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
+ where
+
+ _×-surjection_ : Surjection A B → Surjection C D →
+ Surjection (A ×ₛ C) (B ×ₛ D)
+ A↠B ×-surjection C↠D = record
+ { to = LeftInverse.from inv
+ ; surjective = record
+ { from = LeftInverse.to inv
+ ; right-inverse-of = LeftInverse.left-inverse-of inv
+ }
+ }
+ where
+ open Surjection
+ inv = right-inverse A↠B ×-left-inverse right-inverse C↠D
+
+ _×-inverse_ : Inverse A B → Inverse C D →
+ Inverse (A ×ₛ C) (B ×ₛ D)
+ A↔B ×-inverse C↔D = record
+ { to = Surjection.to surj
+ ; from = Surjection.from surj
+ ; inverse-of = record
+ { left-inverse-of = LeftInverse.left-inverse-of inv
+ ; right-inverse-of = Surjection.right-inverse-of surj
+ }
+ }
+ where
+ open Inverse
+ surj = Inverse.surjection A↔B ×-surjection
+ Inverse.surjection C↔D
+ inv = Inverse.left-inverse A↔B ×-left-inverse
+ Inverse.left-inverse C↔D
+
+------------------------------------------------------------------------
+
+module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
+
+ _×-⇔_ : A ⇔ B → C ⇔ D → (A × C) ⇔ (B × D)
+ _×-⇔_ A⇔B C⇔D =
+ Inverse.equivalence Pointwise-≡↔≡ ⟨∘⟩
+ (A⇔B ×-equivalence C⇔D) ⟨∘⟩
+ Eq.sym (Inverse.equivalence Pointwise-≡↔≡)
+ where open Eq using () renaming (_∘_ to _⟨∘⟩_)
+
+ _×-↣_ : A ↣ B → C ↣ D → (A × C) ↣ (B × D)
+ _×-↣_ A↣B C↣D =
+ Inverse.injection Pointwise-≡↔≡ ⟨∘⟩
+ (A↣B ×-injection C↣D) ⟨∘⟩
+ Inverse.injection (Inv.sym Pointwise-≡↔≡)
+ where open Inj using () renaming (_∘_ to _⟨∘⟩_)
+
+ _×-↞_ : A ↞ B → C ↞ D → (A × C) ↞ (B × D)
+ _×-↞_ A↞B C↞D =
+ Inverse.left-inverse Pointwise-≡↔≡ ⟨∘⟩
+ (A↞B ×-left-inverse C↞D) ⟨∘⟩
+ Inverse.left-inverse (Inv.sym Pointwise-≡↔≡)
+ where open LeftInv using () renaming (_∘_ to _⟨∘⟩_)
+
+ _×-↠_ : A ↠ B → C ↠ D → (A × C) ↠ (B × D)
+ _×-↠_ A↠B C↠D =
+ Inverse.surjection Pointwise-≡↔≡ ⟨∘⟩
+ (A↠B ×-surjection C↠D) ⟨∘⟩
+ Inverse.surjection (Inv.sym Pointwise-≡↔≡)
+ where open Surj using () renaming (_∘_ to _⟨∘⟩_)
+
+ _×-↔_ : A ↔ B → C ↔ D → (A × C) ↔ (B × D)
+ _×-↔_ A↔B C↔D =
+ Pointwise-≡↔≡ ⟨∘⟩
+ (A↔B ×-inverse C↔D) ⟨∘⟩
+ Inv.sym Pointwise-≡↔≡
+ where open Inv using () renaming (_∘_ to _⟨∘⟩_)
+
+_×-cong_ : ∀ {k a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ A ∼[ k ] B → C ∼[ k ] D → (A × C) ∼[ k ] (B × D)
+_×-cong_ {implication} = λ f g → Prod.map f g
+_×-cong_ {reverse-implication} = λ f g → lam (Prod.map (app-← f) (app-← g))
+_×-cong_ {equivalence} = _×-⇔_
+_×-cong_ {injection} = _×-↣_
+_×-cong_ {reverse-injection} = λ f g → lam (app-↢ f ×-↣ app-↢ g)
+_×-cong_ {left-inverse} = _×-↞_
+_×-cong_ {surjection} = _×-↠_
+_×-cong_ {bijection} = _×-↔_
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+infixr 2 _×-Rel_
+_×-Rel_ = Pointwise
+{-# WARNING_ON_USAGE _×-Rel_
+"Warning: _×-Rel_ was deprecated in v0.15.
+Please use Pointwise instead."
+#-}
+Rel↔≡ = Pointwise-≡↔≡
+{-# WARNING_ON_USAGE Rel↔≡
+"Warning: Rel↔≡ was deprecated in v0.15.
+Please use Pointwise-≡↔≡ instead."
+#-}
+_×-reflexive_ = ×-reflexive
+{-# WARNING_ON_USAGE _×-reflexive_
+"Warning: _×-reflexive_ was deprecated in v0.15.
+Please use ×-reflexive instead."
+#-}
+_×-refl_ = ×-refl
+{-# WARNING_ON_USAGE _×-refl_
+"Warning: _×-refl_ was deprecated in v0.15.
+Please use ×-refl instead."
+#-}
+_×-symmetric_ = ×-symmetric
+{-# WARNING_ON_USAGE _×-symmetric_
+"Warning: _×-symmetric_ was deprecated in v0.15.
+Please use ×-symmetric instead."
+#-}
+_×-transitive_ = ×-transitive
+{-# WARNING_ON_USAGE _×-transitive_
+"Warning: _×-transitive_ was deprecated in v0.15.
+Please use ×-transitive instead."
+#-}
+_×-antisymmetric_ = ×-antisymmetric
+{-# WARNING_ON_USAGE _×-antisymmetric_
+"Warning: _×-antisymmetric_ was deprecated in v0.15.
+Please use ×-antisymmetric instead."
+#-}
+_×-≈-respects₂_ = ×-respects₂
+{-# WARNING_ON_USAGE _×-≈-respects₂_
+"Warning: _×-≈-respects₂_ was deprecated in v0.15.
+Please use ×-respects₂ instead."
+#-}
+_×-decidable_ = ×-decidable
+{-# WARNING_ON_USAGE _×-decidable_
+"Warning: _×-decidable_ was deprecated in v0.15.
+Please use ×-decidable instead."
+#-}
+_×-isEquivalence_ = ×-isEquivalence
+{-# WARNING_ON_USAGE _×-isEquivalence_
+"Warning: _×-isEquivalence_ was deprecated in v0.15.
+Please use ×-isEquivalence instead."
+#-}
+_×-isDecEquivalence_ = ×-isDecEquivalence
+{-# WARNING_ON_USAGE _×-isDecEquivalence_
+"Warning: _×-isDecEquivalence_ was deprecated in v0.15.
+Please use ×-isDecEquivalence instead."
+#-}
+_×-isPreorder_ = ×-isPreorder
+{-# WARNING_ON_USAGE _×-isPreorder_
+"Warning: _×-isPreorder_ was deprecated in v0.15.
+Please use ×-isPreorder instead."
+#-}
+_×-isPartialOrder_ = ×-isPartialOrder
+{-# WARNING_ON_USAGE _×-isPartialOrder_
+"Warning: _×-isPartialOrder_ was deprecated in v0.15.
+Please use ×-isPartialOrder instead."
+#-}
+_×-isStrictPartialOrder_ = ×-isStrictPartialOrder
+{-# WARNING_ON_USAGE _×-isStrictPartialOrder_
+"Warning: _×-isStrictPartialOrder_ was deprecated in v0.15.
+Please use ×-isStrictPartialOrder instead."
+#-}
+_×-preorder_ = ×-preorder
+{-# WARNING_ON_USAGE _×-preorder_
+"Warning: _×-preorder_ was deprecated in v0.15.
+Please use ×-preorder instead."
+#-}
+_×-setoid_ = ×-setoid
+{-# WARNING_ON_USAGE _×-setoid_
+"Warning: _×-setoid_ was deprecated in v0.15.
+Please use ×-setoid instead."
+#-}
+_×-decSetoid_ = ×-decSetoid
+{-# WARNING_ON_USAGE _×-decSetoid_
+"Warning: _×-decSetoid_ was deprecated in v0.15.
+Please use ×-decSetoid instead."
+#-}
+_×-poset_ = ×-poset
+{-# WARNING_ON_USAGE _×-poset_
+"Warning: _×-poset_ was deprecated in v0.15.
+Please use ×-poset instead."
+#-}
+_×-strictPartialOrder_ = ×-strictPartialOrder
+{-# WARNING_ON_USAGE _×-strictPartialOrder_
+"Warning: _×-strictPartialOrder_ was deprecated in v0.15.
+Please use ×-strictPartialOrder instead."
+#-}
+_×-≟_ = ≡?×≡?⇒≡?
+{-# WARNING_ON_USAGE _×-≟_
+"Warning: _×-≟_ was deprecated in v0.15.
+Please use ≡?×≡?⇒≡? instead."
+#-}
diff --git a/src/Data/Rational.agda b/src/Data/Rational.agda
index 535d6a0..86e8e27 100644
--- a/src/Data/Rational.agda
+++ b/src/Data/Rational.agda
@@ -12,7 +12,7 @@ open import Function
open import Data.Integer as ℤ using (ℤ; ∣_∣; +_; -_)
open import Data.Integer.Divisibility as ℤDiv using (Coprime)
import Data.Integer.Properties as ℤ
-open import Data.Nat.Divisibility as ℕDiv using (_∣_)
+open import Data.Nat.Divisibility as ℕDiv using (_∣_; ∣-antisym)
import Data.Nat.Coprimality as C
open import Data.Nat as ℕ using (ℕ; zero; suc)
open import Data.Sum
@@ -51,9 +51,9 @@ infixl 7 _÷_
_÷_ : (numerator : ℤ) (denominator : ℕ)
{coprime : True (C.coprime? ∣ numerator ∣ denominator)}
- {≢0 : False (ℕ._≟_ denominator 0)} →
+ {≢0 : False (denominator ℕ.≟ 0)} →
-(n ÷ zero) {≢0 = ()}
+(n ÷ zero) {≢0 = ()}
(n ÷ suc d) {c} = record
{ numerator = n
; denominator-1 = d
@@ -98,8 +98,7 @@ p ≃ q = numerator p ℤ.* denominator q ≡
helper : ∀ n₁ d₁ c₁ n₂ d₂ c₂ →
n₁ ℤ.* + suc d₂ ≡ n₂ ℤ.* + suc d₁ →
(n₁ ÷ suc d₁) {c₁} ≡ (n₂ ÷ suc d₂) {c₂}
- helper n₁ d₁ c₁ n₂ d₂ c₂ eq
- with Poset.antisym ℕDiv.poset 1+d₁∣1+d₂ 1+d₂∣1+d₁
+ helper n₁ d₁ c₁ n₂ d₂ c₂ eq with ∣-antisym 1+d₁∣1+d₂ 1+d₂∣1+d₁
where
1+d₁∣1+d₂ : suc d₁ ∣ suc d₂
1+d₁∣1+d₂ = ℤDiv.coprime-divisor (+ suc d₁) n₁ (+ suc d₂)
@@ -116,10 +115,9 @@ p ≃ q = numerator p ℤ.* denominator q ≡
∣ n₂ ℤ.* + suc d₁ ∣ ≡⟨ cong ∣_∣ (P.sym eq) ⟩
∣ n₁ ℤ.* + suc d₂ ∣ ≡⟨ ℤ.abs-*-commute n₁ (+ suc d₂) ⟩
∣ n₁ ∣ ℕ.* suc d₂ ∎)
-
- helper n₁ d c₁ n₂ .d c₂ eq | refl with ℤ.cancel-*-right
+ helper n₁ d c₁ n₂ .d c₂ eq | refl with ℤ.*-cancelʳ-≡
n₁ n₂ (+ suc d) (λ ()) eq
- helper n d c₁ .n .d c₂ eq | refl | refl with Bool.proof-irrelevance c₁ c₂
+ helper n d c₁ .n .d c₂ eq | refl | refl with Bool.T-irrelevance c₁ c₂
helper n d c .n .d .c eq | refl | refl | refl = refl
------------------------------------------------------------------------
@@ -130,8 +128,8 @@ infix 4 _≟_
_≟_ : Decidable {A = ℚ} _≡_
p ≟ q with ℚ.numerator p ℤ.* ℚ.denominator q ℤ.≟
ℚ.numerator q ℤ.* ℚ.denominator p
-p ≟ q | yes pq≃qp = yes (≃⇒≡ pq≃qp)
-p ≟ q | no ¬pq≃qp = no (¬pq≃qp ∘ ≡⇒≃)
+... | yes pq≃qp = yes (≃⇒≡ pq≃qp)
+... | no ¬pq≃qp = no (¬pq≃qp ∘ ≡⇒≃)
------------------------------------------------------------------------
-- Ordering
diff --git a/src/Data/Rational/Literals.agda b/src/Data/Rational/Literals.agda
new file mode 100644
index 0000000..fb0feea
--- /dev/null
+++ b/src/Data/Rational/Literals.agda
@@ -0,0 +1,35 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Rational Literals
+------------------------------------------------------------------------
+
+module Data.Rational.Literals where
+
+open import Agda.Builtin.FromNat
+open import Agda.Builtin.FromNeg
+open import Data.Unit
+open import Data.Nat
+open import Data.Nat.Coprimality
+open import Data.Integer
+open import Data.Rational
+open import Relation.Nullary.Decidable
+
+fromℤ : ℤ → ℚ
+fromℤ z = record
+ { numerator = z
+ ; denominator-1 = zero
+ ; isCoprime = fromWitness (λ {i} → sym (1-coprimeTo ∣ z ∣))
+ }
+
+number : Number ℚ
+number = record
+ { Constraint = λ _ → ⊤
+ ; fromNat = λ n → fromℤ (+ n)
+ }
+
+negative : Negative ℚ
+negative = record
+ { Constraint = λ _ → ⊤
+ ; fromNeg = λ n → fromℤ (- (+ n))
+ }
diff --git a/src/Data/Rational/Properties.agda b/src/Data/Rational/Properties.agda
index 1f377e9..0ee449b 100644
--- a/src/Data/Rational/Properties.agda
+++ b/src/Data/Rational/Properties.agda
@@ -28,13 +28,13 @@ open import Relation.Binary.PropositionalEquality as P
≤-trans : Transitive _≤_
≤-trans {i = p} {j = q} {k = r} (*≤* le₁) (*≤* le₂)
- = *≤* (ℤₚ.cancel-*-+-right-≤ _ _ _
+ = *≤* (ℤₚ.*-cancelʳ-≤-pos _ _ _
(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₂)))
+ (ℤₚ.*-monoʳ-≤-pos (ℚ.denominator-1 r) le₁)
+ (ℤₚ.*-monoʳ-≤-pos (ℚ.denominator-1 p) le₂)))
where
lemma : ∀ n₁ d₁ n₂ d₂ n₃ d₃ →
n₁ ℤ.* d₂ ℤ.* d₃ ℤ.≤ n₂ ℤ.* d₁ ℤ.* d₃ →
@@ -95,3 +95,5 @@ open import Relation.Binary.PropositionalEquality as P
; isDecTotalOrder = ≤-isDecTotalOrder
}
+≤-irrelevance : Irrelevant _≤_
+≤-irrelevance (*≤* x₁) (*≤* x₂) = P.cong *≤* (ℤₚ.≤-irrelevance x₁ x₂)
diff --git a/src/Data/ReflexiveClosure.agda b/src/Data/ReflexiveClosure.agda
index cf12175..223d171 100644
--- a/src/Data/ReflexiveClosure.agda
+++ b/src/Data/ReflexiveClosure.agda
@@ -2,47 +2,11 @@
-- The Agda standard library
--
-- Reflexive closures
+--
+-- This module is DEPRECATED. Please use the
+-- Relation.Binary.Construct.Closure.Reflexive module directly.
------------------------------------------------------------------------
module Data.ReflexiveClosure where
-open import Data.Unit
-open import Level
-open import Relation.Binary
-open import Relation.Binary.Simple
-
-------------------------------------------------------------------------
--- Reflexive closure
-
-data Refl {a ℓ} {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where
- [_] : ∀ {x y} (x∼y : x ∼ y) → Refl _∼_ x y
- refl : Reflexive (Refl _∼_)
-
--- Map.
-
-map : ∀ {a a′ ℓ ℓ′} {A : Set a} {A′ : Set a′}
- {_R_ : Rel A ℓ} {_R′_ : Rel A′ ℓ′} {f : A → A′} →
- _R_ =[ f ]⇒ _R′_ → Refl _R_ =[ f ]⇒ Refl _R′_
-map R⇒R′ [ xRy ] = [ R⇒R′ xRy ]
-map R⇒R′ refl = refl
-
--- The reflexive closure has no effect on reflexive relations.
-
-drop-refl : ∀ {a ℓ} {A : Set a} {_R_ : Rel A ℓ} →
- Reflexive _R_ → Refl _R_ ⇒ _R_
-drop-refl rfl [ x∼y ] = x∼y
-drop-refl rfl refl = rfl
-
-------------------------------------------------------------------------
--- Example: Maybe
-
-module Maybe where
-
- Maybe : ∀ {ℓ} → Set ℓ → Set ℓ
- Maybe A = Refl (Const A) tt tt
-
- nothing : ∀ {a} {A : Set a} → Maybe A
- nothing = refl
-
- just : ∀ {a} {A : Set a} → A → Maybe A
- just = [_]
+open import Relation.Binary.Construct.Closure.Reflexive public
diff --git a/src/Data/Sign.agda b/src/Data/Sign.agda
index eef2317..199901c 100644
--- a/src/Data/Sign.agda
+++ b/src/Data/Sign.agda
@@ -6,10 +6,9 @@
module Data.Sign where
-open import Relation.Nullary
-open import Relation.Binary
-open import Relation.Binary.Core using (_≡_; refl)
--- Importing Core here ^^^ to keep a small import list
+open import Relation.Binary using (Decidable)
+open import Relation.Binary.PropositionalEquality using (_≡_; refl)
+open import Relation.Nullary using (yes; no)
-- Signs.
diff --git a/src/Data/Sign/Properties.agda b/src/Data/Sign/Properties.agda
index 8ac8bc4..35cc7dc 100644
--- a/src/Data/Sign/Properties.agda
+++ b/src/Data/Sign/Properties.agda
@@ -6,6 +6,8 @@
module Data.Sign.Properties where
+open import Algebra
+open import Algebra.Structures
open import Data.Empty
open import Function
open import Data.Sign
@@ -15,19 +17,21 @@ open import Algebra.FunctionProperties (_≡_ {A = Sign})
-- The opposite of a sign is not equal to the sign.
-opposite-not-equal : ∀ s → s ≢ opposite s
-opposite-not-equal - ()
-opposite-not-equal + ()
+s≢opposite[s] : ∀ s → s ≢ opposite s
+s≢opposite[s] - ()
+s≢opposite[s] + ()
-opposite-cong : ∀ {s t} → opposite s ≡ opposite t → s ≡ t
-opposite-cong { - } { - } refl = refl
-opposite-cong { - } { + } ()
-opposite-cong { + } { - } ()
-opposite-cong { + } { + } refl = refl
+opposite-injective : ∀ {s t} → opposite s ≡ opposite t → s ≡ t
+opposite-injective { - } { - } refl = refl
+opposite-injective { - } { + } ()
+opposite-injective { + } { - } ()
+opposite-injective { + } { + } refl = refl
------------------------------------------------------------------------
-- _*_
+-- Algebraic properties of _*_
+
*-identityˡ : LeftIdentity + _*_
*-identityˡ _ = refl
@@ -51,19 +55,83 @@ opposite-cong { + } { + } refl = 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ʳ-≡ : RightCancellative _*_
+*-cancelʳ-≡ - - _ = refl
+*-cancelʳ-≡ - + eq = ⊥-elim (s≢opposite[s] _ $ sym eq)
+*-cancelʳ-≡ + - eq = ⊥-elim (s≢opposite[s] _ eq)
+*-cancelʳ-≡ + + _ = refl
+
+*-cancelˡ-≡ : LeftCancellative _*_
+*-cancelˡ-≡ - eq = opposite-injective eq
+*-cancelˡ-≡ + eq = eq
+
+*-cancel-≡ : Cancellative _*_
+*-cancel-≡ = *-cancelˡ-≡ , *-cancelʳ-≡
+
+*-isSemigroup : IsSemigroup _≡_ _*_
+*-isSemigroup = record
+ { isEquivalence = isEquivalence
+ ; assoc = *-assoc
+ ; ∙-cong = cong₂ _*_
+ }
+
+*-semigroup : Semigroup _ _
+*-semigroup = record
+ { Carrier = Sign
+ ; _≈_ = _≡_
+ ; _∙_ = _*_
+ ; isSemigroup = *-isSemigroup
+ }
-cancel-*-left : LeftCancellative _*_
-cancel-*-left - eq = opposite-cong eq
-cancel-*-left + eq = eq
+*-isMonoid : IsMonoid _≡_ _*_ +
+*-isMonoid = record
+ { isSemigroup = *-isSemigroup
+ ; identity = *-identity
+ }
-*-cancellative : Cancellative _*_
-*-cancellative = cancel-*-left , cancel-*-right
+*-monoid : Monoid _ _
+*-monoid = record
+ { Carrier = Sign
+ ; _≈_ = _≡_
+ ; _∙_ = _*_
+ ; ε = +
+ ; isMonoid = *-isMonoid
+ }
+
+-- Other properties of _*_
s*s≡+ : ∀ s → s * s ≡ +
s*s≡+ + = refl
s*s≡+ - = refl
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+opposite-not-equal = s≢opposite[s]
+{-# WARNING_ON_USAGE opposite-not-equal
+"Warning: opposite-not-equal was deprecated in v0.15.
+Please use s≢opposite[s] instead."
+#-}
+opposite-cong = opposite-injective
+{-# WARNING_ON_USAGE opposite-cong
+"Warning: opposite-cong was deprecated in v0.15.
+Please use opposite-injective instead."
+#-}
+cancel-*-left = *-cancelˡ-≡
+{-# WARNING_ON_USAGE cancel-*-left
+"Warning: cancel-*-left was deprecated in v0.15.
+Please use *-cancelˡ-≡ instead."
+#-}
+cancel-*-right = *-cancelʳ-≡
+{-# WARNING_ON_USAGE cancel-*-right
+"Warning: cancel-*-right was deprecated in v0.15.
+Please use *-cancelʳ-≡ instead."
+#-}
+*-cancellative = *-cancel-≡
+{-# WARNING_ON_USAGE *-cancellative
+"Warning: *-cancellative was deprecated in v0.15.
+Please use *-cancel-≡ instead."
+#-}
diff --git a/src/Data/Star.agda b/src/Data/Star.agda
index b85e442..00f1f5a 100644
--- a/src/Data/Star.agda
+++ b/src/Data/Star.agda
@@ -2,137 +2,11 @@
-- The Agda standard library
--
-- The reflexive transitive closures of McBride, Norell and Jansson
+--
+-- This module is DEPRECATED. Please use the
+-- Relation.Binary.Construct.Closure.ReflexiveTransitive module directly
------------------------------------------------------------------------
--- This module could be placed under Relation.Binary. However, since
--- its primary purpose is to be used for _data_ it has been placed
--- under Data instead.
-
module Data.Star where
-open import Relation.Binary
-open import Function
-open import Level
-
-infixr 5 _◅_
-
--- Reflexive transitive closure.
-
-data Star {i t} {I : Set i} (T : Rel I t) : Rel I (i ⊔ t) where
- ε : Reflexive (Star T)
- _◅_ : ∀ {i j k} (x : T i j) (xs : Star T j k) → Star T i k
- -- The type of _◅_ is Trans T (Star T) (Star T); I expanded
- -- the definition in order to be able to name the arguments (x
- -- and xs).
-
--- Append/transitivity.
-
-infixr 5 _◅◅_
-
-_◅◅_ : ∀ {i t} {I : Set i} {T : Rel I t} → Transitive (Star T)
-ε ◅◅ ys = ys
-(x ◅ xs) ◅◅ ys = x ◅ (xs ◅◅ ys)
-
--- Sometimes you want to view cons-lists as snoc-lists. Then the
--- following "constructor" is handy. Note that this is _not_ snoc for
--- cons-lists, it is just a synonym for cons (with a different
--- argument order).
-
-infixl 5 _▻_
-
-_▻_ : ∀ {i t} {I : Set i} {T : Rel I t} {i j k} →
- Star T j k → T i j → Star T i k
-_▻_ = flip _◅_
-
--- A corresponding variant of append.
-
-infixr 5 _▻▻_
-
-_▻▻_ : ∀ {i t} {I : Set i} {T : Rel I t} {i j k} →
- Star T j k → Star T i j → Star T i k
-_▻▻_ = flip _◅◅_
-
--- A generalised variant of map which allows the index type to change.
-
-gmap : ∀ {i j t u} {I : Set i} {T : Rel I t} {J : Set j} {U : Rel J u} →
- (f : I → J) → T =[ f ]⇒ U → Star T =[ f ]⇒ Star U
-gmap f g ε = ε
-gmap f g (x ◅ xs) = g x ◅ gmap f g xs
-
-map : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} →
- T ⇒ U → Star T ⇒ Star U
-map = gmap id
-
--- A generalised variant of fold.
-
-gfold : ∀ {i j t p} {I : Set i} {J : Set j} {T : Rel I t}
- (f : I → J) (P : Rel J p) →
- Trans T (P on f) (P on f) →
- TransFlip (Star T) (P on f) (P on f)
-gfold f P _⊕_ ∅ ε = ∅
-gfold f P _⊕_ ∅ (x ◅ xs) = x ⊕ gfold f P _⊕_ ∅ xs
-
-fold : ∀ {i t p} {I : Set i} {T : Rel I t} (P : Rel I p) →
- Trans T P P → Reflexive P → Star T ⇒ P
-fold P _⊕_ ∅ = gfold id P _⊕_ ∅
-
-gfoldl : ∀ {i j t p} {I : Set i} {J : Set j} {T : Rel I t}
- (f : I → J) (P : Rel J p) →
- Trans (P on f) T (P on f) →
- Trans (P on f) (Star T) (P on f)
-gfoldl f P _⊕_ ∅ ε = ∅
-gfoldl f P _⊕_ ∅ (x ◅ xs) = gfoldl f P _⊕_ (∅ ⊕ x) xs
-
-foldl : ∀ {i t p} {I : Set i} {T : Rel I t} (P : Rel I p) →
- Trans P T P → Reflexive P → Star T ⇒ P
-foldl P _⊕_ ∅ = gfoldl id P _⊕_ ∅
-
-concat : ∀ {i t} {I : Set i} {T : Rel I t} → Star (Star T) ⇒ Star T
-concat {T = T} = fold (Star T) _◅◅_ ε
-
--- If the underlying relation is symmetric, then the reflexive
--- transitive closure is also symmetric.
-
-revApp : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} →
- Sym T U → ∀ {i j k} → Star T j i → Star U j k → Star U i k
-revApp rev ε ys = ys
-revApp rev (x ◅ xs) ys = revApp rev xs (rev x ◅ ys)
-
-reverse : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} →
- Sym T U → Sym (Star T) (Star U)
-reverse rev xs = revApp rev xs ε
-
--- Reflexive transitive closures form a (generalised) monad.
-
--- return could also be called singleton.
-
-return : ∀ {i t} {I : Set i} {T : Rel I t} → T ⇒ Star T
-return x = x ◅ ε
-
--- A generalised variant of the Kleisli star (flip bind, or
--- concatMap).
-
-kleisliStar : ∀ {i j t u}
- {I : Set i} {J : Set j} {T : Rel I t} {U : Rel J u}
- (f : I → J) → T =[ f ]⇒ Star U → Star T =[ f ]⇒ Star U
-kleisliStar f g = concat ∘′ gmap f g
-
-_⋆ : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} →
- T ⇒ Star U → Star T ⇒ Star U
-_⋆ = kleisliStar id
-
-infixl 1 _>>=_
-
-_>>=_ : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} {i j} →
- Star T i j → T ⇒ Star U → Star U i j
-m >>= f = (f ⋆) m
-
--- Note that the monad-like structure above is not an indexed monad
--- (as defined in Category.Monad.Indexed). If it were, then _>>=_
--- would have a type similar to
---
--- ∀ {I} {T U : Rel I t} {i j k} →
--- Star T i j → (T i j → Star U j k) → Star U i k.
--- ^^^^^
--- Note, however, that there is no scope for applying T to any indices
--- in the definition used in Category.Monad.Indexed.
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive public
diff --git a/src/Data/Star/BoundedVec.agda b/src/Data/Star/BoundedVec.agda
index cef4e6c..376fe14 100644
--- a/src/Data/Star/BoundedVec.agda
+++ b/src/Data/Star/BoundedVec.agda
@@ -8,16 +8,16 @@
module Data.Star.BoundedVec where
-open import Data.Star
+import Data.Maybe.Base as Maybe
open import Data.Star.Nat
open import Data.Star.Decoration
open import Data.Star.Pointer
open import Data.Star.List using (List)
open import Data.Unit
open import Function
-import Data.Maybe.Base as Maybe
open import Relation.Binary
open import Relation.Binary.Consequences
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
------------------------------------------------------------------------
-- The type
diff --git a/src/Data/Star/Decoration.agda b/src/Data/Star/Decoration.agda
index 0e954cc..fb98b19 100644
--- a/src/Data/Star/Decoration.agda
+++ b/src/Data/Star/Decoration.agda
@@ -6,49 +6,50 @@
module Data.Star.Decoration where
-open import Data.Star
-open import Relation.Binary
-open import Function
open import Data.Unit
+open import Function
open import Level
+open import Relation.Binary
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
-- A predicate on relation "edges" (think of the relation as a graph).
-EdgePred : {I : Set} → Rel I zero → Set₁
-EdgePred T = ∀ {i j} → T i j → Set
+EdgePred : {ℓ r : Level} (p : Level) {I : Set ℓ} → Rel I r → Set (suc p ⊔ ℓ ⊔ r)
+EdgePred p T = ∀ {i j} → T i j → Set p
+
-data NonEmptyEdgePred {I : Set}
- (T : Rel I zero) (P : EdgePred T) : Set where
+data NonEmptyEdgePred {ℓ r p : Level} {I : Set ℓ} (T : Rel I r)
+ (P : EdgePred p T) : Set (ℓ ⊔ r ⊔ p) where
nonEmptyEdgePred : ∀ {i j} {x : T i j}
(p : P x) → NonEmptyEdgePred T P
-- Decorating an edge with more information.
-data DecoratedWith {I : Set} {T : Rel I zero} (P : EdgePred T)
- : Rel (NonEmpty (Star T)) zero where
+data DecoratedWith {ℓ r p : Level} {I : Set ℓ} {T : Rel I r} (P : EdgePred p T)
+ : Rel (NonEmpty (Star T)) p where
↦ : ∀ {i j k} {x : T i j} {xs : Star T j k}
(p : P x) → DecoratedWith P (nonEmpty (x ◅ xs)) (nonEmpty xs)
-edge : ∀ {I} {T : Rel I zero} {P : EdgePred T} {i j} →
- DecoratedWith {T = T} P i j → NonEmpty T
-edge (↦ {x = x} p) = nonEmpty x
+module _ {ℓ r p : Level} {I : Set ℓ} {T : Rel I r} {P : EdgePred p T} where
+
+ edge : ∀ {i j} → DecoratedWith {T = T} P i j → NonEmpty T
+ edge (↦ {x = x} p) = nonEmpty x
-decoration : ∀ {I} {T : Rel I zero} {P : EdgePred T} {i j} →
- (d : DecoratedWith {T = T} P i j) →
- P (NonEmpty.proof (edge d))
-decoration (↦ p) = p
+ decoration : ∀ {i j} → (d : DecoratedWith {T = T} P i j) →
+ P (NonEmpty.proof (edge d))
+ decoration (↦ p) = p
-- Star-lists decorated with extra information. All P xs means that
-- all edges in xs satisfy P.
-All : ∀ {I} {T : Rel I zero} → EdgePred T → EdgePred (Star T)
+All : ∀ {ℓ r p} {I : Set ℓ} {T : Rel I r} → EdgePred p T → EdgePred (ℓ ⊔ (r ⊔ p)) (Star T)
All P {j = j} xs =
Star (DecoratedWith P) (nonEmpty xs) (nonEmpty {y = j} ε)
-- We can map over decorated vectors.
-gmapAll : ∀ {I} {T : Rel I zero} {P : EdgePred T}
- {J} {U : Rel J zero} {Q : EdgePred U}
+gmapAll : ∀ {ℓ ℓ′ r p q} {I : Set ℓ} {T : Rel I r} {P : EdgePred p T}
+ {J : Set ℓ′} {U : Rel J r} {Q : EdgePred q U}
{i j} {xs : Star T i j}
(f : I → J) (g : T =[ f ]⇒ U) →
(∀ {i j} {x : T i j} → P x → Q (g x)) →
@@ -59,7 +60,8 @@ gmapAll f g h (↦ x ◅ xs) = ↦ (h x) ◅ gmapAll f g h xs
-- Since we don't automatically have gmap id id xs ≡ xs it is easier
-- to implement mapAll in terms of map than in terms of gmapAll.
-mapAll : ∀ {I} {T : Rel I zero} {P Q : EdgePred T} {i j} {xs : Star T i j} →
+mapAll : ∀ {ℓ r p q} {I : Set ℓ} {T : Rel I r}
+ {P : EdgePred p T} {Q : EdgePred q T} {i j} {xs : Star T i j} →
(∀ {i j} {x : T i j} → P x → Q x) →
All P xs → All Q xs
mapAll {P = P} {Q} f ps = map F ps
@@ -69,7 +71,7 @@ mapAll {P = P} {Q} f ps = map F ps
-- We can decorate star-lists with universally true predicates.
-decorate : ∀ {I} {T : Rel I zero} {P : EdgePred T} {i j} →
+decorate : ∀ {ℓ r p} {I : Set ℓ} {T : Rel I r} {P : EdgePred p T} {i j} →
(∀ {i j} (x : T i j) → P x) →
(xs : Star T i j) → All P xs
decorate f ε = ε
@@ -79,13 +81,13 @@ decorate f (x ◅ xs) = ↦ (f x) ◅ decorate f xs
infixr 5 _◅◅◅_ _▻▻▻_
-_◅◅◅_ : ∀ {I} {T : Rel I zero} {P : EdgePred T}
+_◅◅◅_ : ∀ {ℓ r p} {I : Set ℓ} {T : Rel I r} {P : EdgePred p T}
{i j k} {xs : Star T i j} {ys : Star T j k} →
All P xs → All P ys → All P (xs ◅◅ ys)
ε ◅◅◅ ys = ys
(↦ x ◅ xs) ◅◅◅ ys = ↦ x ◅ xs ◅◅◅ ys
-_▻▻▻_ : ∀ {I} {T : Rel I zero} {P : EdgePred T}
+_▻▻▻_ : ∀ {ℓ r p} {I : Set ℓ} {T : Rel I r} {P : EdgePred p T}
{i j k} {xs : Star T j k} {ys : Star T i j} →
All P xs → All P ys → All P (xs ▻▻ ys)
_▻▻▻_ = flip _◅◅◅_
diff --git a/src/Data/Star/Environment.agda b/src/Data/Star/Environment.agda
index 826d8e7..fe906f0 100644
--- a/src/Data/Star/Environment.agda
+++ b/src/Data/Star/Environment.agda
@@ -4,36 +4,38 @@
-- Environments (heterogeneous collections)
------------------------------------------------------------------------
-module Data.Star.Environment (Ty : Set) where
+module Data.Star.Environment {ℓ} (Ty : Set ℓ) where
-open import Data.Star
+open import Level
open import Data.Star.List
open import Data.Star.Decoration
open import Data.Star.Pointer as Pointer hiding (lookup)
open import Data.Unit
+open import Function hiding (_∋_)
open import Relation.Binary.PropositionalEquality
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
-- Contexts, listing the types of all the elements in an environment.
-Ctxt : Set
+Ctxt : Set ℓ
Ctxt = List Ty
-- Variables (de Bruijn indices); pointers into environments.
infix 4 _∋_
-_∋_ : Ctxt → Ty → Set
-Γ ∋ σ = Any (λ _ → ⊤) (_≡_ σ) Γ
+_∋_ : Ctxt → Ty → Set ℓ
+Γ ∋ σ = Any (const (Lift ℓ ⊤)) (σ ≡_) Γ
vz : ∀ {Γ σ} → Γ ▻ σ ∋ σ
vz = this refl
vs : ∀ {Γ σ τ} → Γ ∋ τ → Γ ▻ σ ∋ τ
-vs = that tt
+vs = that _
-- Environments. The T function maps types to element types.
-Env : (Ty → Set) → Ctxt → Set
+Env : ∀ {e} → (Ty → Set e) → (Ctxt → Set (ℓ ⊔ e))
Env T Γ = All T Γ
-- A safe lookup function for environments.
diff --git a/src/Data/Star/Fin.agda b/src/Data/Star/Fin.agda
index add8e50..e20fc41 100644
--- a/src/Data/Star/Fin.agda
+++ b/src/Data/Star/Fin.agda
@@ -1,12 +1,11 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Finite sets defined in terms of Data.Star
+-- Finite sets defined using the reflexive-transitive closure, Star
------------------------------------------------------------------------
module Data.Star.Fin where
-open import Data.Star
open import Data.Star.Nat as ℕ using (ℕ)
open import Data.Star.Pointer
open import Data.Unit
diff --git a/src/Data/Star/List.agda b/src/Data/Star/List.agda
index 54aeb1a..9365f86 100644
--- a/src/Data/Star/List.agda
+++ b/src/Data/Star/List.agda
@@ -1,29 +1,30 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Lists defined in terms of Data.Star
+-- Lists defined in terms of the reflexive-transitive closure, Star
------------------------------------------------------------------------
module Data.Star.List where
-open import Data.Star
-open import Data.Unit
-open import Relation.Binary.Simple
open import Data.Star.Nat
+open import Data.Unit
+open import Relation.Binary.Construct.Always using (Always)
+open import Relation.Binary.Construct.Constant using (Const)
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
-- Lists.
-List : Set → Set
-List a = Star (Const a) tt tt
+List : ∀ {a} → Set a → Set a
+List A = Star (Const A) tt tt
-- Nil and cons.
-[] : ∀ {a} → List a
+[] : ∀ {a} {A : Set a} → List A
[] = ε
infixr 5 _∷_
-_∷_ : ∀ {a} → a → List a → List a
+_∷_ : ∀ {a} {A : Set a} → A → List A → List A
_∷_ = _◅_
-- The sum of the elements in a list containing natural numbers.
diff --git a/src/Data/Star/Nat.agda b/src/Data/Star/Nat.agda
index a54a465..7597032 100644
--- a/src/Data/Star/Nat.agda
+++ b/src/Data/Star/Nat.agda
@@ -1,16 +1,16 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Natural numbers defined in terms of Data.Star
+-- Natural numbers defined using the reflexive-transitive closure, Star
------------------------------------------------------------------------
module Data.Star.Nat where
-open import Data.Star
open import Data.Unit
open import Function
open import Relation.Binary
-open import Relation.Binary.Simple
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
+open import Relation.Binary.Construct.Always using (Always)
-- Natural numbers.
diff --git a/src/Data/Star/Pointer.agda b/src/Data/Star/Pointer.agda
index 52d9737..83983c6 100644
--- a/src/Data/Star/Pointer.agda
+++ b/src/Data/Star/Pointer.agda
@@ -4,21 +4,22 @@
-- Pointers into star-lists
------------------------------------------------------------------------
-module Data.Star.Pointer where
+module Data.Star.Pointer {ℓ} {I : Set ℓ} where
-open import Data.Star
-open import Data.Star.Decoration
-open import Relation.Binary
open import Data.Maybe.Base using (Maybe; nothing; just)
+open import Data.Star.Decoration
open import Data.Unit
open import Function
open import Level
+open import Relation.Binary
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
-- Pointers into star-lists. The edge pointed to is decorated with Q,
-- while other edges are decorated with P.
-data Pointer {I : Set} {T : Rel I zero} (P Q : EdgePred T)
- : Rel (Maybe (NonEmpty (Star T))) zero where
+data Pointer {r p q} {T : Rel I r}
+ (P : EdgePred p T) (Q : EdgePred q T)
+ : Rel (Maybe (NonEmpty (Star T))) (p ⊔ q) where
step : ∀ {i j k} {x : T i j} {xs : Star T j k}
(p : P x) → Pointer P Q (just (nonEmpty (x ◅ xs)))
(just (nonEmpty xs))
@@ -30,61 +31,60 @@ data Pointer {I : Set} {T : Rel I zero} (P Q : EdgePred T)
-- is basically a prefix of xs; the existence of such a prefix
-- guarantees that xs is non-empty.
-Any : ∀ {I} {T : Rel I zero} (P Q : EdgePred T) → EdgePred (Star T)
+Any : ∀ {r p q} {T : Rel I r} (P : EdgePred p T) (Q : EdgePred q T) →
+ EdgePred (ℓ ⊔ (r ⊔ (p ⊔ q))) (Star T)
Any P Q xs = Star (Pointer P Q) (just (nonEmpty xs)) nothing
-this : ∀ {I} {T : Rel I zero} {P Q : EdgePred T}
- {i j k} {x : T i j} {xs : Star T j k} →
- Q x → Any P Q (x ◅ xs)
-this q = done q ◅ ε
+module _ {r p q} {T : Rel I r} {P : EdgePred p T} {Q : EdgePred q T} where
+
+ this : ∀ {i j k} {x : T i j} {xs : Star T j k} →
+ Q x → Any P Q (x ◅ xs)
+ this q = done q ◅ ε
-that : ∀ {I} {T : Rel I zero} {P Q : EdgePred T}
- {i j k} {x : T i j} {xs : Star T j k} →
- P x → Any P Q xs → Any P Q (x ◅ xs)
-that p = _◅_ (step p)
+ that : ∀ {i j k} {x : T i j} {xs : Star T j k} →
+ P x → Any P Q xs → Any P Q (x ◅ xs)
+ that p = _◅_ (step p)
-- Safe lookup.
-data Result {I : Set} (T : Rel I zero) (P Q : EdgePred T) : Set where
- result : ∀ {i j} {x : T i j}
- (p : P x) (q : Q x) → Result T P Q
+data Result {r p q} (T : Rel I r)
+ (P : EdgePred p T) (Q : EdgePred q T) : Set (ℓ ⊔ r ⊔ p ⊔ q) where
+ result : ∀ {i j} {x : T i j} (p : P x) (q : Q x) → Result T P Q
-- The first argument points out which edge to extract. The edge is
-- returned, together with proofs that it satisfies Q and R.
-lookup : ∀ {I} {T : Rel I zero} {P Q R : EdgePred T}
- {i j} {xs : Star T i j} →
- Any P Q xs → All R xs → Result T Q R
-lookup (done q ◅ ε) (↦ r ◅ _) = result q r
-lookup (step p ◅ ps) (↦ r ◅ rs) = lookup ps rs
-lookup (done _ ◅ () ◅ _) _
+module _ {t p q} {T : Rel I t} {P : EdgePred p T} {Q : EdgePred q T} where
+
+ lookup : ∀ {r} {R : EdgePred r T} {i j} {xs : Star T i j} →
+ Any P Q xs → All R xs → Result T Q R
+ lookup (done q ◅ ε) (↦ r ◅ _) = result q r
+ lookup (step p ◅ ps) (↦ r ◅ rs) = lookup ps rs
+ lookup (done _ ◅ () ◅ _) _
-- We can define something resembling init.
-prefixIndex : ∀ {I} {T : Rel I zero} {P Q : EdgePred T}
- {i j} {xs : Star T i j} →
- Any P Q xs → I
-prefixIndex (done {i = i} q ◅ _) = i
-prefixIndex (step p ◅ ps) = prefixIndex ps
+ prefixIndex : ∀ {i j} {xs : Star T i j} → Any P Q xs → I
+ prefixIndex (done {i = i} q ◅ _) = i
+ prefixIndex (step p ◅ ps) = prefixIndex ps
-prefix : ∀ {I} {T : Rel I zero} {P Q : EdgePred T} {i j} {xs : Star T i j} →
- (ps : Any P Q xs) → Star T i (prefixIndex ps)
-prefix (done q ◅ _) = ε
-prefix (step {x = x} p ◅ ps) = x ◅ prefix ps
+ prefix : ∀ {i j} {xs : Star T i j} →
+ (ps : Any P Q xs) → Star T i (prefixIndex ps)
+ prefix (done q ◅ _) = ε
+ prefix (step {x = x} p ◅ ps) = x ◅ prefix ps
-- Here we are taking the initial segment of ps (all elements but the
-- last, i.e. all edges satisfying P).
-init : ∀ {I} {T : Rel I zero} {P Q : EdgePred T} {i j} {xs : Star T i j} →
- (ps : Any P Q xs) → All P (prefix ps)
-init (done q ◅ _) = ε
-init (step p ◅ ps) = ↦ p ◅ init ps
+ init : ∀ {i j} {xs : Star T i j} →
+ (ps : Any P Q xs) → All P (prefix ps)
+ init (done q ◅ _) = ε
+ init (step p ◅ ps) = ↦ p ◅ init ps
-- One can simplify the implementation by not carrying around the
-- indices in the type:
-last : ∀ {I} {T : Rel I zero} {P Q : EdgePred T}
- {i j} {xs : Star T i j} →
- Any P Q xs → NonEmptyEdgePred T Q
-last ps with lookup ps (decorate (const tt) _)
-... | result q _ = nonEmptyEdgePred q
+ last : ∀ {i j} {xs : Star T i j} →
+ Any P Q xs → NonEmptyEdgePred T Q
+ last ps with lookup {r = p} ps (decorate (const (lift tt)) _)
+ ... | result q _ = nonEmptyEdgePred q
diff --git a/src/Data/Star/Properties.agda b/src/Data/Star/Properties.agda
index a5d88bf..e3672bb 100644
--- a/src/Data/Star/Properties.agda
+++ b/src/Data/Star/Properties.agda
@@ -2,94 +2,13 @@
-- The Agda standard library
--
-- Some properties related to Data.Star
+--
+-- This module is DEPRECATED. Please use the
+-- Relation.Binary.Construct.Closure.ReflexiveTransitive.Properties
+-- module directly.
------------------------------------------------------------------------
module Data.Star.Properties where
-open import Data.Star
-open import Function
-open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as PropEq
- using (_≡_; refl; sym; cong; cong₂)
-import Relation.Binary.PreorderReasoning as PreR
-
-◅◅-assoc : ∀ {i t} {I : Set i} {T : Rel I t} {i j k l}
- (xs : Star T i j) (ys : Star T j k)
- (zs : Star T k l) →
- (xs ◅◅ ys) ◅◅ zs ≡ xs ◅◅ (ys ◅◅ zs)
-◅◅-assoc ε ys zs = refl
-◅◅-assoc (x ◅ xs) ys zs = cong (_◅_ x) (◅◅-assoc xs ys zs)
-
-gmap-id : ∀ {i t} {I : Set i} {T : Rel I t} {i j} (xs : Star T i j) →
- gmap id id xs ≡ xs
-gmap-id ε = refl
-gmap-id (x ◅ xs) = cong (_◅_ x) (gmap-id xs)
-
-gmap-∘ : ∀ {i t} {I : Set i} {T : Rel I t}
- {j u} {J : Set j} {U : Rel J u}
- {k v} {K : Set k} {V : Rel K v}
- (f : J → K) (g : U =[ f ]⇒ V)
- (f′ : I → J) (g′ : T =[ f′ ]⇒ U)
- {i j} (xs : Star T i j) →
- (gmap {U = V} f g ∘ gmap f′ g′) xs ≡ gmap (f ∘ f′) (g ∘ g′) xs
-gmap-∘ f g f′ g′ ε = refl
-gmap-∘ f g f′ g′ (x ◅ xs) = cong (_◅_ (g (g′ x))) (gmap-∘ f g f′ g′ xs)
-
-gmap-◅◅ : ∀ {i t j u}
- {I : Set i} {T : Rel I t} {J : Set j} {U : Rel J u}
- (f : I → J) (g : T =[ f ]⇒ U)
- {i j k} (xs : Star T i j) (ys : Star T j k) →
- gmap {U = U} f g (xs ◅◅ ys) ≡ gmap f g xs ◅◅ gmap f g ys
-gmap-◅◅ f g ε ys = refl
-gmap-◅◅ f g (x ◅ xs) ys = cong (_◅_ (g x)) (gmap-◅◅ f g xs ys)
-
-gmap-cong : ∀ {i t j u}
- {I : Set i} {T : Rel I t} {J : Set j} {U : Rel J u}
- (f : I → J) (g : T =[ f ]⇒ U) (g′ : T =[ f ]⇒ U) →
- (∀ {i j} (x : T i j) → g x ≡ g′ x) →
- ∀ {i j} (xs : Star T i j) →
- gmap {U = U} f g xs ≡ gmap f g′ xs
-gmap-cong f g g′ eq ε = refl
-gmap-cong f g g′ eq (x ◅ xs) = cong₂ _◅_ (eq x) (gmap-cong f g g′ eq xs)
-
-fold-◅◅ : ∀ {i p} {I : Set i}
- (P : Rel I p) (_⊕_ : Transitive P) (∅ : Reflexive P) →
- (∀ {i j} (x : P i j) → (∅ ⊕ x) ≡ x) →
- (∀ {i j k l} (x : P i j) (y : P j k) (z : P k l) →
- ((x ⊕ y) ⊕ z) ≡ (x ⊕ (y ⊕ z))) →
- ∀ {i j k} (xs : Star P i j) (ys : Star P j k) →
- fold P _⊕_ ∅ (xs ◅◅ ys) ≡ (fold P _⊕_ ∅ xs ⊕ fold P _⊕_ ∅ ys)
-fold-◅◅ P _⊕_ ∅ left-unit assoc ε ys = sym (left-unit _)
-fold-◅◅ P _⊕_ ∅ left-unit assoc (x ◅ xs) ys = begin
- (x ⊕ fold P _⊕_ ∅ (xs ◅◅ ys)) ≡⟨ cong (_⊕_ x) $
- fold-◅◅ P _⊕_ ∅ left-unit assoc xs ys ⟩
- (x ⊕ (fold P _⊕_ ∅ xs ⊕ fold P _⊕_ ∅ ys)) ≡⟨ sym (assoc x _ _) ⟩
- ((x ⊕ fold P _⊕_ ∅ xs) ⊕ fold P _⊕_ ∅ ys) ∎
- where open PropEq.≡-Reasoning
-
--- Reflexive transitive closures are preorders.
-
-preorder : ∀ {i t} {I : Set i} (T : Rel I t) → Preorder _ _ _
-preorder T = record
- { _≈_ = _≡_
- ; _∼_ = Star T
- ; isPreorder = record
- { isEquivalence = PropEq.isEquivalence
- ; reflexive = reflexive
- ; trans = _◅◅_
- }
- }
- where
- reflexive : _≡_ ⇒ Star T
- reflexive refl = ε
-
--- Preorder reasoning for Star.
-
-module StarReasoning {i t} {I : Set i} (T : Rel I t) where
- open PreR (preorder T) public
- renaming (_∼⟨_⟩_ to _⟶⋆⟨_⟩_; _≈⟨_⟩_ to _≡⟨_⟩_)
-
- infixr 2 _⟶⟨_⟩_
-
- _⟶⟨_⟩_ : ∀ x {y z} → T x y → y IsRelatedTo z → x IsRelatedTo z
- x ⟶⟨ x⟶y ⟩ y⟶⋆z = x ⟶⋆⟨ x⟶y ◅ ε ⟩ y⟶⋆z
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive.Properties
+ public
diff --git a/src/Data/Star/Vec.agda b/src/Data/Star/Vec.agda
index e7e0a53..e10b941 100644
--- a/src/Data/Star/Vec.agda
+++ b/src/Data/Star/Vec.agda
@@ -1,12 +1,11 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Vectors defined in terms of Data.Star
+-- Vectors defined in terms of the reflexive-transitive closure, Star
------------------------------------------------------------------------
module Data.Star.Vec where
-open import Data.Star
open import Data.Star.Nat
open import Data.Star.Fin using (Fin)
open import Data.Star.Decoration
@@ -14,6 +13,7 @@ open import Data.Star.Pointer as Pointer hiding (lookup)
open import Data.Star.List using (List)
open import Relation.Binary
open import Relation.Binary.Consequences
+open import Relation.Binary.Construct.Closure.ReflexiveTransitive
open import Function
open import Data.Unit
@@ -21,45 +21,45 @@ open import Data.Unit
-- information (i.e. elements).
Vec : Set → ℕ → Set
-Vec a = All (λ _ → a)
+Vec A = All (λ _ → A)
-- Nil and cons.
-[] : ∀ {a} → Vec a zero
+[] : ∀ {A} → Vec A zero
[] = ε
infixr 5 _∷_
-_∷_ : ∀ {a n} → a → Vec a n → Vec a (suc n)
+_∷_ : ∀ {A n} → A → Vec A n → Vec A (suc n)
x ∷ xs = ↦ x ◅ xs
-- Projections.
-head : ∀ {a n} → Vec a (1# + n) → a
+head : ∀ {A n} → Vec A (1# + n) → A
head (↦ x ◅ _) = x
-tail : ∀ {a n} → Vec a (1# + n) → Vec a n
+tail : ∀ {A n} → Vec A (1# + n) → Vec A n
tail (↦ _ ◅ xs) = xs
-- Append.
infixr 5 _++_
-_++_ : ∀ {a m n} → Vec a m → Vec a n → Vec a (m + n)
+_++_ : ∀ {A m n} → Vec A m → Vec A n → Vec A (m + n)
_++_ = _◅◅◅_
-- Safe lookup.
-lookup : ∀ {a n} → Fin n → Vec a n → a
+lookup : ∀ {A n} → Fin n → Vec A n → A
lookup i xs with Pointer.lookup i xs
... | result _ x = x
------------------------------------------------------------------------
-- Conversions
-fromList : ∀ {a} → (xs : List a) → Vec a (length xs)
+fromList : ∀ {A} → (xs : List A) → Vec A (length xs)
fromList ε = []
fromList (x ◅ xs) = x ∷ fromList xs
-toList : ∀ {a n} → Vec a n → List a
+toList : ∀ {A n} → Vec A n → List A
toList = gmap (const tt) decoration
diff --git a/src/Data/Stream.agda b/src/Data/Stream.agda
deleted file mode 100644
index 3ae4e28..0000000
--- a/src/Data/Stream.agda
+++ /dev/null
@@ -1,162 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- Streams
-------------------------------------------------------------------------
-
-module Data.Stream where
-
-open import Coinduction
-open import Data.Colist using (Colist; []; _∷_)
-open import Data.Vec using (Vec; []; _∷_)
-open import Data.Nat.Base using (ℕ; zero; suc)
-open import Relation.Binary
-open import Relation.Binary.PropositionalEquality as P using (_≡_)
-
-------------------------------------------------------------------------
--- The type
-
-infixr 5 _∷_
-
-data Stream {a} (A : Set a) : Set a where
- _∷_ : (x : A) (xs : ∞ (Stream A)) → Stream A
-
-{-# 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
-
-head : ∀ {a} {A : Set a} → Stream A → A
-head (x ∷ xs) = x
-
-tail : ∀ {a} {A : Set a} → Stream A → Stream A
-tail (x ∷ xs) = ♭ xs
-
-map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → Stream A → Stream B
-map f (x ∷ xs) = f x ∷ ♯ map f (♭ xs)
-
-zipWith : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
- (A → B → C) → Stream A → Stream B → Stream C
-zipWith _∙_ (x ∷ xs) (y ∷ ys) = (x ∙ y) ∷ ♯ zipWith _∙_ (♭ xs) (♭ ys)
-
-take : ∀ {a} {A : Set a} n → Stream A → Vec A n
-take zero xs = []
-take (suc n) (x ∷ xs) = x ∷ take n (♭ xs)
-
-drop : ∀ {a} {A : Set a} → ℕ → Stream A → Stream A
-drop zero xs = xs
-drop (suc n) (x ∷ xs) = drop n (♭ xs)
-
-repeat : ∀ {a} {A : Set a} → A → Stream A
-repeat x = x ∷ ♯ repeat x
-
-iterate : ∀ {a} {A : Set a} → (A → A) → A → Stream A
-iterate f x = x ∷ ♯ iterate f (f x)
-
--- Interleaves the two streams.
-
-infixr 5 _⋎_
-
-_⋎_ : ∀ {a} {A : Set a} → Stream A → Stream A → Stream A
-(x ∷ xs) ⋎ ys = x ∷ ♯ (ys ⋎ ♭ xs)
-
-mutual
-
- -- Takes every other element from the stream, starting with the
- -- first one.
-
- evens : ∀ {a} {A : Set a} → Stream A → Stream A
- evens (x ∷ xs) = x ∷ ♯ odds (♭ xs)
-
- -- Takes every other element from the stream, starting with the
- -- second one.
-
- odds : ∀ {a} {A : Set a} → Stream A → Stream A
- odds (x ∷ xs) = evens (♭ xs)
-
-toColist : ∀ {a} {A : Set a} → Stream A → Colist A
-toColist (x ∷ xs) = x ∷ ♯ toColist (♭ xs)
-
-lookup : ∀ {a} {A : Set a} → ℕ → Stream A → A
-lookup zero (x ∷ xs) = x
-lookup (suc n) (x ∷ xs) = lookup n (♭ xs)
-
-infixr 5 _++_
-
-_++_ : ∀ {a} {A : Set a} → Colist A → Stream A → Stream A
-[] ++ ys = ys
-(x ∷ xs) ++ ys = x ∷ ♯ (♭ xs ++ ys)
-
-------------------------------------------------------------------------
--- Equality and other relations
-
--- xs ≈ ys means that xs and ys are equal.
-
-infix 4 _≈_
-
-data _≈_ {a} {A : Set a} : Stream A → Stream A → Set a where
- _∷_ : ∀ {x y xs ys}
- (x≡ : x ≡ y) (xs≈ : ∞ (♭ xs ≈ ♭ ys)) → x ∷ xs ≈ y ∷ ys
-
--- x ∈ xs means that x is a member of xs.
-
-infix 4 _∈_
-
-data _∈_ {a} {A : Set a} : A → Stream A → Set a where
- here : ∀ {x xs} → x ∈ x ∷ xs
- there : ∀ {x y xs} (x∈xs : x ∈ ♭ xs) → x ∈ y ∷ xs
-
--- xs ⊑ ys means that xs is a prefix of ys.
-
-infix 4 _⊑_
-
-data _⊑_ {a} {A : Set a} : Colist A → Stream A → Set a where
- [] : ∀ {ys} → [] ⊑ ys
- _∷_ : ∀ x {xs ys} (p : ∞ (♭ xs ⊑ ♭ ys)) → x ∷ xs ⊑ x ∷ ys
-
-------------------------------------------------------------------------
--- Some proofs
-
-setoid : ∀ {a} → Set a → Setoid _ _
-setoid A = record
- { Carrier = Stream A
- ; _≈_ = _≈_ {A = A}
- ; isEquivalence = record
- { refl = refl
- ; sym = sym
- ; trans = trans
- }
- }
- where
- refl : Reflexive _≈_
- refl {_ ∷ _} = P.refl ∷ ♯ refl
-
- sym : Symmetric _≈_
- sym (x≡ ∷ xs≈) = P.sym x≡ ∷ ♯ sym (♭ xs≈)
-
- trans : Transitive _≈_
- trans (x≡ ∷ xs≈) (y≡ ∷ ys≈) = P.trans x≡ y≡ ∷ ♯ trans (♭ xs≈) (♭ ys≈)
-
-head-cong : ∀ {a} {A : Set a} {xs ys : Stream A} → xs ≈ ys → head xs ≡ head ys
-head-cong (x≡ ∷ _) = x≡
-
-tail-cong : ∀ {a} {A : Set a} {xs ys : Stream A} → xs ≈ ys → tail xs ≈ tail ys
-tail-cong (_ ∷ xs≈) = ♭ xs≈
-
-map-cong : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) {xs ys} →
- xs ≈ ys → map f xs ≈ map f ys
-map-cong f (x≡ ∷ xs≈) = P.cong f x≡ ∷ ♯ map-cong f (♭ xs≈)
-
-zipWith-cong : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}
- (_∙_ : A → B → C) {xs xs′ ys ys′} →
- xs ≈ xs′ → ys ≈ ys′ →
- zipWith _∙_ xs ys ≈ zipWith _∙_ xs′ ys′
-zipWith-cong _∙_ (x≡ ∷ xs≈) (y≡ ∷ ys≈) =
- P.cong₂ _∙_ x≡ y≡ ∷ ♯ zipWith-cong _∙_ (♭ xs≈) (♭ ys≈)
-
-infixr 5 _⋎-cong_
-
-_⋎-cong_ : ∀ {a} {A : Set a} {xs xs′ ys ys′ : Stream A} →
- xs ≈ xs′ → ys ≈ ys′ → xs ⋎ ys ≈ xs′ ⋎ ys′
-(x ∷ xs≈) ⋎-cong ys≈ = x ∷ ♯ (ys≈ ⋎-cong ♭ xs≈)
diff --git a/src/Data/String.agda b/src/Data/String.agda
index c8198d3..22b1ab0 100644
--- a/src/Data/String.agda
+++ b/src/Data/String.agda
@@ -6,76 +6,32 @@
module Data.String where
-open import Data.List.Base as List using (_∷_; []; List)
open import Data.Vec as Vec using (Vec)
-open import Data.Colist as Colist using (Colist)
open import Data.Char as Char using (Char)
-open import Data.Bool.Base using (Bool; true; false)
-open import Function
-open import Relation.Nullary
-open import Relation.Nullary.Decidable
-open import Relation.Binary
-open import Relation.Binary.List.StrictLex as StrictLex
-import Relation.Binary.On as On
-open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
-open import Relation.Binary.PropositionalEquality.TrustMe
+open import Relation.Binary using (Setoid; StrictTotalOrder)
+open import Data.List.Relation.Lex.Strict as StrictLex
+import Relation.Binary.Construct.On as On
+import Relation.Binary.PropositionalEquality as PropEq
-open import Data.String.Base public
-
--- Possibly infinite strings.
+------------------------------------------------------------------------
+-- Re-export contents of base publically
-Costring : Set
-Costring = Colist Char
+open import Data.String.Base public
------------------------------------------------------------------------
-- Operations
-toVec : (s : String) → Vec Char (List.length (toList s))
+toVec : (s : String) → Vec Char (length s)
toVec s = Vec.fromList (toList s)
-toCostring : String → Costring
-toCostring = Colist.fromList ∘ toList
-
--- Informative equality test.
-
-infix 4 _≟_
-
-_≟_ : Decidable {A = String} _≡_
-s₁ ≟ s₂ with primStringEquality s₁ s₂
-... | true = yes trustMe
-... | false = no whatever
- where postulate whatever : _
-
--- Boolean equality test.
---
--- Why is the definition _==_ = primStringEquality not used? One
--- reason is that the present definition can sometimes improve type
--- inference, at least with the version of Agda that is current at the
--- time of writing: see unit-test below.
-
-infix 4 _==_
-
-_==_ : String → String → Bool
-s₁ == s₂ = ⌊ s₁ ≟ s₂ ⌋
-
-private
-
- -- The following unit test does not type-check (at the time of
- -- writing) if _==_ is replaced by primStringEquality.
-
- data P : (String → Bool) → Set where
- p : (c : String) → P (_==_ c)
-
- unit-test : P (_==_ "")
- unit-test = p _
+------------------------------------------------------------------------
+-- Equality over strings
setoid : Setoid _ _
setoid = PropEq.setoid String
-decSetoid : DecSetoid _ _
-decSetoid = PropEq.decSetoid _≟_
-
--- Lexicographic ordering of strings.
+------------------------------------------------------------------------
+-- A lexicographic ordering on strings.
strictTotalOrder : StrictTotalOrder _ _ _
strictTotalOrder =
diff --git a/src/Data/String/Base.agda b/src/Data/String/Base.agda
index 90b9a43..df29823 100644
--- a/src/Data/String/Base.agda
+++ b/src/Data/String/Base.agda
@@ -6,30 +6,30 @@
module Data.String.Base where
-open import Data.List.Base as List using (_∷_; []; List)
-open import Data.Bool.Base using (Bool)
-open import Data.Char.Core using (Char)
-open import Relation.Binary.Core using (_≡_)
-open import Relation.Binary.PropositionalEquality.TrustMe using (trustMe)
+open import Data.Nat.Base as Nat using (ℕ)
+open import Data.List.Base as List using (List)
+open import Data.List.NonEmpty as NE using (List⁺)
+open import Agda.Builtin.Char using (Char)
+open import Function
+open import Relation.Binary.PropositionalEquality using (_≡_)
------------------------------------------------------------------------
-- From Agda.Builtin
open import Agda.Builtin.String public
- using ( String
- ; primStringAppend
- ; primStringToList
- ; primStringFromList
- ; primStringEquality
- ; primShowString )
+ using
+ ( String
+ ; primStringAppend
+ ; primStringToList
+ ; primStringFromList
+ ; primStringEquality
+ ; primShowString
+ )
------------------------------------------------------------------------
-- Operations
-infixr 5 _++_
-
-_++_ : String → String → String
-_++_ = primStringAppend
+-- Conversion functions
toList : String → List Char
toList = primStringToList
@@ -37,15 +37,28 @@ toList = primStringToList
fromList : List Char → String
fromList = primStringFromList
-toList∘fromList : ∀ s → toList (fromList s) ≡ s
-toList∘fromList s = trustMe
+fromList⁺ : List⁺ Char → String
+fromList⁺ = fromList ∘ NE.toList
-fromList∘toList : ∀ s → fromList (toList s) ≡ s
-fromList∘toList s = trustMe
+-- List-like functions
-unlines : List String → String
-unlines [] = ""
-unlines (x ∷ xs) = x ++ "\n" ++ unlines xs
+infixr 5 _++_
+_++_ : String → String → String
+_++_ = primStringAppend
+
+length : String → ℕ
+length = List.length ∘ toList
+
+replicate : ℕ → Char → String
+replicate n = fromList ∘ List.replicate n
+
+concat : List String → String
+concat = List.foldr _++_ ""
+
+-- String-specific functions
show : String → String
show = primShowString
+
+unlines : List String → String
+unlines = concat ∘ List.intersperse "\n"
diff --git a/src/Data/String/Literals.agda b/src/Data/String/Literals.agda
new file mode 100644
index 0000000..33d4fb5
--- /dev/null
+++ b/src/Data/String/Literals.agda
@@ -0,0 +1,17 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- String Literals
+------------------------------------------------------------------------
+
+module Data.String.Literals where
+
+open import Agda.Builtin.FromString
+open import Data.Unit
+open import Agda.Builtin.String
+
+isString : IsString String
+isString = record
+ { Constraint = λ _ → ⊤
+ ; fromString = λ s → s
+ }
diff --git a/src/Data/String/Unsafe.agda b/src/Data/String/Unsafe.agda
new file mode 100644
index 0000000..3f0581a
--- /dev/null
+++ b/src/Data/String/Unsafe.agda
@@ -0,0 +1,65 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Unsafe String operations and proofs
+------------------------------------------------------------------------
+
+module Data.String.Unsafe where
+
+open import Data.String
+open import Data.Bool.Base using (Bool; true; false)
+open import Relation.Binary using (Decidable; DecSetoid)
+open import Relation.Binary.PropositionalEquality as PropEq using (_≡_)
+open import Relation.Binary.PropositionalEquality.TrustMe using (trustMe)
+open import Relation.Nullary using (yes; no)
+open import Relation.Nullary.Decidable using (⌊_⌋)
+
+------------------------------------------------------------------------
+-- An informative equality test.
+
+infix 4 _≟_
+
+_≟_ : Decidable {A = String} _≡_
+s₁ ≟ s₂ with primStringEquality s₁ s₂
+... | true = yes trustMe
+... | false = no whatever
+ where postulate whatever : _
+
+------------------------------------------------------------------------
+-- Boolean equality test.
+--
+-- Why is the definition _==_ = primStringEquality not used? One
+-- reason is that the present definition can sometimes improve type
+-- inference, at least with the version of Agda that is current at the
+-- time of writing: see unit-test below.
+
+infix 4 _==_
+
+_==_ : String → String → Bool
+s₁ == s₂ = ⌊ s₁ ≟ s₂ ⌋
+
+private
+
+ -- The following unit test does not type-check (at the time of
+ -- writing) if _==_ is replaced by primStringEquality.
+
+ data P : (String → Bool) → Set where
+ p : (c : String) → P (_==_ c)
+
+ unit-test : P (_==_ "")
+ unit-test = p _
+
+------------------------------------------------------------------------
+-- Equality
+
+decSetoid : DecSetoid _ _
+decSetoid = PropEq.decSetoid _≟_
+
+------------------------------------------------------------------------
+-- Other properties
+
+toList∘fromList : ∀ s → toList (fromList s) ≡ s
+toList∘fromList s = trustMe
+
+fromList∘toList : ∀ s → fromList (toList s) ≡ s
+fromList∘toList s = trustMe
diff --git a/src/Data/Sum.agda b/src/Data/Sum.agda
index 2ea1c07..cf59e6c 100644
--- a/src/Data/Sum.agda
+++ b/src/Data/Sum.agda
@@ -10,62 +10,38 @@ open import Function
open import Data.Unit.Base using (⊤; tt)
open import Data.Maybe.Base using (Maybe; just; nothing)
open import Level
+open import Agda.Builtin.Equality
------------------------------------------------------------------------
--- Definition
+-- Re-export content from base module
-infixr 1 _⊎_
-
-data _⊎_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
- inj₁ : (x : A) → A ⊎ B
- inj₂ : (y : B) → A ⊎ B
-
-{-# FOREIGN GHC type AgdaEither a b c d = Either c d #-}
-{-# COMPILE GHC _⊎_ = data MAlonzo.Code.Data.Sum.AgdaEither (Left | Right) #-}
+open import Data.Sum.Base public
------------------------------------------------------------------------
--- Functions
-
-[_,_] : ∀ {a b c} {A : Set a} {B : Set b} {C : A ⊎ B → Set c} →
- ((x : A) → C (inj₁ x)) → ((x : B) → C (inj₂ x)) →
- ((x : A ⊎ B) → C x)
-[ f , g ] (inj₁ x) = f x
-[ f , g ] (inj₂ y) = g y
-
-[_,_]′ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
- (A → C) → (B → C) → (A ⊎ B → C)
-[_,_]′ = [_,_]
-
-map : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
- (A → C) → (B → D) → (A ⊎ B → C ⊎ D)
-map f g = [ inj₁ ∘ f , inj₂ ∘ g ]
-
-infixr 1 _-⊎-_
+-- Additional functions
-_-⊎-_ : ∀ {a b c d} {A : Set a} {B : Set b} →
- (A → B → Set c) → (A → B → Set d) → (A → B → Set (c ⊔ d))
-f -⊎- g = f -[ _⊎_ ]- g
+module _ {a b} {A : Set a} {B : Set b} where
-isInj₁ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → Maybe A
-isInj₁ (inj₁ x) = just x
-isInj₁ (inj₂ y) = nothing
+ isInj₁ : A ⊎ B → Maybe A
+ isInj₁ (inj₁ x) = just x
+ 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
+ isInj₂ : 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 → Set a
+ From-inj₁ (inj₁ _) = A
+ From-inj₁ (inj₂ _) = Lift a ⊤
-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₁ : (x : A ⊎ B) → From-inj₁ x
+ from-inj₁ (inj₁ x) = x
+ from-inj₁ (inj₂ _) = _
-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 → Set b
+ From-inj₂ (inj₁ _) = Lift b ⊤
+ From-inj₂ (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
+ from-inj₂ : (x : A ⊎ B) → From-inj₂ x
+ from-inj₂ (inj₁ _) = _
+ from-inj₂ (inj₂ x) = x
diff --git a/src/Data/Sum/Base.agda b/src/Data/Sum/Base.agda
new file mode 100644
index 0000000..270478f
--- /dev/null
+++ b/src/Data/Sum/Base.agda
@@ -0,0 +1,56 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Sums (disjoint unions)
+------------------------------------------------------------------------
+
+module Data.Sum.Base where
+
+open import Function using (_∘_; _-[_]-_ ; id)
+open import Level using (_⊔_)
+
+------------------------------------------------------------------------
+-- Definition
+
+infixr 1 _⊎_
+
+data _⊎_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
+ inj₁ : (x : A) → A ⊎ B
+ inj₂ : (y : B) → A ⊎ B
+
+{-# FOREIGN GHC type AgdaEither a b c d = Either c d #-}
+{-# COMPILE GHC _⊎_ = data AgdaEither (Left | Right) #-}
+
+------------------------------------------------------------------------
+-- Functions
+
+[_,_] : ∀ {a b c} {A : Set a} {B : Set b} {C : A ⊎ B → Set c} →
+ ((x : A) → C (inj₁ x)) → ((x : B) → C (inj₂ x)) →
+ ((x : A ⊎ B) → C x)
+[ f , g ] (inj₁ x) = f x
+[ f , g ] (inj₂ y) = g y
+
+[_,_]′ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
+ (A → C) → (B → C) → (A ⊎ B → C)
+[_,_]′ = [_,_]
+
+swap : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → B ⊎ A
+swap (inj₁ x) = inj₂ x
+swap (inj₂ x) = inj₁ x
+
+map : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ (A → C) → (B → D) → (A ⊎ B → C ⊎ D)
+map f g = [ inj₁ ∘ f , inj₂ ∘ g ]
+
+map₁ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c}→
+ (A → C) → (A ⊎ B → C ⊎ B)
+map₁ f = map f id
+
+map₂ : ∀ {a b d} {A : Set a} {B : Set b} {D : Set d} →
+ (B → D) → (A ⊎ B → A ⊎ D)
+map₂ = map id
+
+infixr 1 _-⊎-_
+_-⊎-_ : ∀ {a b c d} {A : Set a} {B : Set b} →
+ (A → B → Set c) → (A → B → Set d) → (A → B → Set (c ⊔ d))
+f -⊎- g = f -[ _⊎_ ]- g
diff --git a/src/Data/Sum/Categorical/Examples.agda b/src/Data/Sum/Categorical/Examples.agda
new file mode 100644
index 0000000..b1d0ee0
--- /dev/null
+++ b/src/Data/Sum/Categorical/Examples.agda
@@ -0,0 +1,52 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Usage examples of the categorical view of the Sum type
+------------------------------------------------------------------------
+
+module Data.Sum.Categorical.Examples where
+
+open import Level
+open import Data.Sum
+import Data.Sum.Categorical.Left as Sumₗ
+open import Category.Functor
+open import Category.Monad
+
+-- Note that these examples are simple unit tests, because the type
+-- checker verifies them.
+
+private
+ module Examplesₗ {a b} {A : Set a} {B : Set b} where
+
+ open import Agda.Builtin.Equality
+ open import Function
+ module Sₗ = Sumₗ A b
+
+ open RawFunctor Sₗ.functor
+
+ -- This type to the right of ⊎ needs to be a "lifted" version of (B : Set b)
+ -- that lives in the universe (Set (a ⊔ b)).
+ fmapId : (x : A ⊎ (Lift a B)) → (id <$> x) ≡ x
+ fmapId (inj₁ x) = refl
+ fmapId (inj₂ y) = refl
+
+
+ open RawMonad Sₗ.monad
+
+ -- Now, let's show that "return" is a unit for >>=. We use Lift in exactly
+ -- the same way as above. The data (x : B) then needs to be "lifted" to
+ -- this new type (Lift B).
+ returnUnitL : ∀ {x : B} {f : Lift a B → A ⊎ (Lift a B)}
+ → ((return (lift x)) >>= f) ≡ f (lift x)
+ returnUnitL = refl
+
+ returnUnitR : (x : A ⊎ (Lift a B)) → (x >>= return) ≡ x
+ returnUnitR (inj₁ _) = refl
+ returnUnitR (inj₂ _) = refl
+
+ -- And another (limited version of a) monad law...
+ bindCompose : ∀ {f g : Lift a B → A ⊎ (Lift a B)}
+ → (x : A ⊎ (Lift a B))
+ → ((x >>= f) >>= g) ≡ (x >>= (λ y → (f y >>= g)))
+ bindCompose (inj₁ x) = refl
+ bindCompose (inj₂ y) = refl
diff --git a/src/Data/Sum/Categorical/Left.agda b/src/Data/Sum/Categorical/Left.agda
new file mode 100644
index 0000000..eed1db2
--- /dev/null
+++ b/src/Data/Sum/Categorical/Left.agda
@@ -0,0 +1,76 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A Categorical view of the Sum type (Left-biased)
+------------------------------------------------------------------------
+
+open import Level
+
+module Data.Sum.Categorical.Left {a} (A : Set a) (b : Level) where
+
+open import Data.Sum
+open import Category.Functor
+open import Category.Applicative
+open import Category.Monad
+import Function.Identity.Categorical as Id
+open import Function
+
+-- To minimize the universe level of the RawFunctor, we require that elements of
+-- B are "lifted" to a copy of B at a higher universe level (a ⊔ b). See the
+-- examples for how this is done.
+
+------------------------------------------------------------------------
+-- Left-biased monad instance for _⊎_
+
+Sumₗ : Set (a ⊔ b) → Set (a ⊔ b)
+Sumₗ B = A ⊎ B
+
+functor : RawFunctor Sumₗ
+functor = record { _<$>_ = map₂ }
+
+applicative : RawApplicative Sumₗ
+applicative = record
+ { pure = inj₂
+ ; _⊛_ = [ const ∘ inj₁ , map₂ ]′
+ }
+
+-- The monad instance also requires some mucking about with universe levels.
+monadT : RawMonadT (_∘′ Sumₗ)
+monadT M = record
+ { return = M.pure ∘ inj₂
+ ; _>>=_ = λ ma f → ma M.>>= [ M.pure ∘′ inj₁ , f ]′
+ } where module M = RawMonad M
+
+monad : RawMonad Sumₗ
+monad = monadT Id.monad
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {F} (App : RawApplicative {a ⊔ b} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → Sumₗ (F A) → F (Sumₗ A)
+ sequenceA (inj₁ a) = pure (inj₁ a)
+ sequenceA (inj₂ x) = inj₂ <$> x
+
+ mapA : ∀ {A B} → (A → F B) → Sumₗ A → F (Sumₗ B)
+ mapA f = sequenceA ∘ map₂ f
+
+ forA : ∀ {A B} → Sumₗ A → (A → F B) → F (Sumₗ B)
+ forA = flip mapA
+
+module _ {M} (Mon : RawMonad {a ⊔ b} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM : ∀ {A} → Sumₗ (M A) → M (Sumₗ A)
+ sequenceM = sequenceA App
+
+ mapM : ∀ {A B} → (A → M B) → Sumₗ A → M (Sumₗ B)
+ mapM = mapA App
+
+ forM : ∀ {A B} → Sumₗ A → (A → M B) → M (Sumₗ B)
+ forM = forA App
+
diff --git a/src/Data/Sum/Categorical/Right.agda b/src/Data/Sum/Categorical/Right.agda
new file mode 100644
index 0000000..cdd0404
--- /dev/null
+++ b/src/Data/Sum/Categorical/Right.agda
@@ -0,0 +1,68 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A Categorical view of the Sum type (Right-biased)
+------------------------------------------------------------------------
+
+open import Level
+
+module Data.Sum.Categorical.Right (a : Level) {b} (B : Set b) where
+
+open import Data.Sum
+open import Category.Functor
+open import Category.Applicative
+open import Category.Monad
+open import Function
+import Function.Identity.Categorical as Id
+
+Sumᵣ : Set (a ⊔ b) → Set (a ⊔ b)
+Sumᵣ A = A ⊎ B
+
+functor : RawFunctor Sumᵣ
+functor = record { _<$>_ = map₁ }
+
+applicative : RawApplicative Sumᵣ
+applicative = record
+ { pure = inj₁
+ ; _⊛_ = [ map₁ , const ∘ inj₂ ]′
+ }
+
+monadT : RawMonadT (_∘′ Sumᵣ)
+monadT M = record
+ { return = M.pure ∘′ inj₁
+ ; _>>=_ = λ ma f → ma M.>>= [ f , M.pure ∘′ inj₂ ]′
+ } where module M = RawMonad M
+
+monad : RawMonad Sumᵣ
+monad = monadT Id.monad
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {F} (App : RawApplicative {a ⊔ b} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → Sumᵣ (F A) → F (Sumᵣ A)
+ sequenceA (inj₂ a) = pure (inj₂ a)
+ sequenceA (inj₁ x) = inj₁ <$> x
+
+ mapA : ∀ {A B} → (A → F B) → Sumᵣ A → F (Sumᵣ B)
+ mapA f = sequenceA ∘ map₁ f
+
+ forA : ∀ {A B} → Sumᵣ A → (A → F B) → F (Sumᵣ B)
+ forA = flip mapA
+
+module _ {M} (Mon : RawMonad {a ⊔ b} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM : ∀ {A} → Sumᵣ (M A) → M (Sumᵣ A)
+ sequenceM = sequenceA App
+
+ mapM : ∀ {A B} → (A → M B) → Sumᵣ A → M (Sumᵣ B)
+ mapM = mapA App
+
+ forM : ∀ {A B} → Sumᵣ A → (A → M B) → M (Sumᵣ B)
+ forM = forA App
+
diff --git a/src/Data/Sum/Properties.agda b/src/Data/Sum/Properties.agda
new file mode 100644
index 0000000..c0df120
--- /dev/null
+++ b/src/Data/Sum/Properties.agda
@@ -0,0 +1,22 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of sums (disjoint unions)
+------------------------------------------------------------------------
+
+module Data.Sum.Properties where
+
+open import Data.Sum
+open import Function
+open import Relation.Binary.PropositionalEquality
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ inj₁-injective : ∀ {x y} → (A ⊎ B ∋ inj₁ x) ≡ inj₁ y → x ≡ y
+ inj₁-injective refl = refl
+
+ inj₂-injective : ∀ {x y} → (A ⊎ B ∋ inj₂ x) ≡ inj₂ y → x ≡ y
+ inj₂-injective refl = refl
+
+ swap-involutive : swap {A = A} {B} ∘ swap ≗ id
+ swap-involutive = [ (λ _ → refl) , (λ _ → refl) ]
diff --git a/src/Data/Sum/Relation/Core.agda b/src/Data/Sum/Relation/Core.agda
new file mode 100644
index 0000000..82f2521
--- /dev/null
+++ b/src/Data/Sum/Relation/Core.agda
@@ -0,0 +1,132 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Sums of binary relations
+------------------------------------------------------------------------
+
+module Data.Sum.Relation.Core where
+
+open import Data.Sum using (_⊎_; inj₁; inj₂)
+open import Data.Product using (_,_; proj₁; proj₂)
+open import Data.Unit.Base using (⊤)
+open import Data.Empty using (⊥)
+open import Function using (_⟨_⟩_; _∘_; flip)
+open import Level using (_⊔_)
+open import Relation.Nullary using (Dec; yes; no)
+open import Relation.Binary
+
+module _ {a₁ a₂} {A₁ : Set a₁} {A₂ : Set a₂} where
+
+ ----------------------------------------------------------------------
+ -- Generalised sum of relations
+
+ data ⊎ʳ {ℓ₁ ℓ₂} (P : Set) (_∼₁_ : Rel A₁ ℓ₁) (_∼₂_ : Rel A₂ ℓ₂) :
+ A₁ ⊎ A₂ → A₁ ⊎ A₂ → Set (a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂) where
+ ₁∼₂ : ∀ {x y} (p : P) → ⊎ʳ P _∼₁_ _∼₂_ (inj₁ x) (inj₂ y)
+ ₁∼₁ : ∀ {x y} (x∼₁y : x ∼₁ y) → ⊎ʳ P _∼₁_ _∼₂_ (inj₁ x) (inj₁ y)
+ ₂∼₂ : ∀ {x y} (x∼₂y : x ∼₂ y) → ⊎ʳ P _∼₁_ _∼₂_ (inj₂ x) (inj₂ y)
+
+ ----------------------------------------------------------------------
+ -- Helpers
+
+ module _ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} where
+
+ drop-inj₁ : ∀ {P x y} → inj₁ x ⟨ ⊎ʳ P ∼₁ ∼₂ ⟩ inj₁ y → ∼₁ x y
+ drop-inj₁ (₁∼₁ x∼y) = x∼y
+
+ drop-inj₂ : ∀ {P x y} → inj₂ x ⟨ ⊎ʳ P ∼₁ ∼₂ ⟩ inj₂ y → ∼₂ x y
+ drop-inj₂ (₂∼₂ x∼y) = x∼y
+
+ ----------------------------------------------------------------------
+ -- Some properties which are preserved by ⊎ʳ
+
+ module _ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} where
+
+ _⊎-refl_ : Reflexive ∼₁ → Reflexive ∼₂ →
+ ∀ {P} → Reflexive (⊎ʳ P ∼₁ ∼₂)
+ refl₁ ⊎-refl refl₂ = refl
+ where
+ refl : Reflexive (⊎ʳ _ _ _)
+ refl {x = inj₁ _} = ₁∼₁ refl₁
+ refl {x = inj₂ _} = ₂∼₂ refl₂
+
+ _⊎-transitive_ : Transitive ∼₁ → Transitive ∼₂ →
+ ∀ {P} → Transitive (⊎ʳ P ∼₁ ∼₂)
+ trans₁ ⊎-transitive trans₂ = trans
+ where
+ trans : Transitive (⊎ʳ _ _ _)
+ trans (₁∼₁ x∼y) (₁∼₁ y∼z) = ₁∼₁ (trans₁ x∼y y∼z)
+ trans (₂∼₂ x∼y) (₂∼₂ y∼z) = ₂∼₂ (trans₂ x∼y y∼z)
+ trans (₁∼₂ p) (₂∼₂ _) = ₁∼₂ p
+ trans (₁∼₁ _) (₁∼₂ p) = ₁∼₂ p
+
+
+ _⊎-asymmetric_ : Asymmetric ∼₁ → Asymmetric ∼₂ →
+ ∀ {P} → Asymmetric (⊎ʳ P ∼₁ ∼₂)
+ asym₁ ⊎-asymmetric asym₂ = asym
+ where
+ asym : Asymmetric (⊎ʳ _ _ _)
+ asym (₁∼₁ x<y) (₁∼₁ y<x) = asym₁ x<y y<x
+ asym (₂∼₂ x<y) (₂∼₂ y<x) = asym₂ x<y y<x
+ asym (₁∼₂ _) ()
+
+ ⊎-decidable : Decidable ∼₁ → Decidable ∼₂ →
+ ∀ {P} → (∀ {x y} → Dec (inj₁ x ⟨ ⊎ʳ P ∼₁ ∼₂ ⟩ inj₂ y)) →
+ Decidable (⊎ʳ P ∼₁ ∼₂)
+ ⊎-decidable dec₁ dec₂ {P} dec₁₂ = dec
+ where
+ dec : Decidable (⊎ʳ P ∼₁ ∼₂)
+ dec (inj₁ x) (inj₂ y) = dec₁₂
+ dec (inj₂ x) (inj₁ y) = no (λ())
+ dec (inj₁ x) (inj₁ y) with dec₁ x y
+ ... | yes x∼y = yes (₁∼₁ x∼y)
+ ... | no x≁y = no (x≁y ∘ drop-inj₁)
+ dec (inj₂ x) (inj₂ y) with dec₂ x y
+ ... | yes x∼y = yes (₂∼₂ x∼y)
+ ... | no x≁y = no (x≁y ∘ drop-inj₂)
+
+ module _ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {≈₁ : Rel A₁ ℓ₂}
+ {ℓ₃ ℓ₄} {∼₂ : Rel A₂ ℓ₃} {≈₂ : Rel A₂ ℓ₄} where
+
+ _⊎-reflexive_ : ≈₁ ⇒ ∼₁ → ≈₂ ⇒ ∼₂ →
+ ∀ {P} → (⊎ʳ ⊥ ≈₁ ≈₂) ⇒ (⊎ʳ P ∼₁ ∼₂)
+ refl₁ ⊎-reflexive refl₂ = refl
+ where
+ refl : (⊎ʳ _ _ _ ) ⇒ (⊎ʳ _ _ _)
+ refl (₁∼₁ x₁≈y₁) = ₁∼₁ (refl₁ x₁≈y₁)
+ refl (₂∼₂ x₂≈y₂) = ₂∼₂ (refl₂ x₂≈y₂)
+ refl (₁∼₂ ())
+
+ _⊎-irreflexive_ : Irreflexive ≈₁ ∼₁ → Irreflexive ≈₂ ∼₂ →
+ ∀ {P} → Irreflexive (⊎ʳ ⊥ ≈₁ ≈₂) (⊎ʳ P ∼₁ ∼₂)
+ irrefl₁ ⊎-irreflexive irrefl₂ = irrefl
+ where
+ irrefl : Irreflexive (⊎ʳ ⊥ _ _) (⊎ʳ _ _ _)
+ irrefl (₁∼₁ x₁≈y₁) (₁∼₁ x₁<y₁) = irrefl₁ x₁≈y₁ x₁<y₁
+ irrefl (₂∼₂ x₂≈y₂) (₂∼₂ x₂<y₂) = irrefl₂ x₂≈y₂ x₂<y₂
+ irrefl (₁∼₂ ()) _
+
+ _⊎-antisymmetric_ : Antisymmetric ≈₁ ∼₁ → Antisymmetric ≈₂ ∼₂ →
+ ∀ {P} → Antisymmetric (⊎ʳ ⊥ ≈₁ ≈₂) (⊎ʳ P ∼₁ ∼₂)
+ antisym₁ ⊎-antisymmetric antisym₂ = antisym
+ where
+ antisym : Antisymmetric (⊎ʳ ⊥ _ _) (⊎ʳ _ _ _)
+ antisym (₁∼₁ x≤y) (₁∼₁ y≤x) = ₁∼₁ (antisym₁ x≤y y≤x)
+ antisym (₂∼₂ x≤y) (₂∼₂ y≤x) = ₂∼₂ (antisym₂ x≤y y≤x)
+ antisym (₁∼₂ _) ()
+
+ _⊎-≈-respects₂_ : ∼₁ Respects₂ ≈₁ → ∼₂ Respects₂ ≈₂ →
+ ∀ {P} → (⊎ʳ P ∼₁ ∼₂) Respects₂ (⊎ʳ ⊥ ≈₁ ≈₂)
+ _⊎-≈-respects₂_ resp₁ resp₂ {P} = resp¹ , resp²
+ where
+ resp¹ : ∀ {x} → ((⊎ʳ P ∼₁ ∼₂) x) Respects (⊎ʳ ⊥ ≈₁ ≈₂)
+ resp¹ (₁∼₁ y≈y') (₁∼₁ x∼y) = ₁∼₁ (proj₁ resp₁ y≈y' x∼y)
+ resp¹ (₂∼₂ y≈y') (₂∼₂ x∼y) = ₂∼₂ (proj₁ resp₂ y≈y' x∼y)
+ resp¹ (₂∼₂ y≈y') (₁∼₂ p) = (₁∼₂ p)
+ resp¹ (₁∼₂ ()) _
+
+ resp² : ∀ {y} → (flip (⊎ʳ P ∼₁ ∼₂) y) Respects (⊎ʳ ⊥ ≈₁ ≈₂)
+ resp² (₁∼₁ x≈x') (₁∼₁ x∼y) = ₁∼₁ (proj₂ resp₁ x≈x' x∼y)
+ resp² (₂∼₂ x≈x') (₂∼₂ x∼y) = ₂∼₂ (proj₂ resp₂ x≈x' x∼y)
+ resp² (₁∼₁ x≈x') (₁∼₂ p) = (₁∼₂ p)
+ resp² (₁∼₂ ()) _
diff --git a/src/Data/Sum/Relation/LeftOrder.agda b/src/Data/Sum/Relation/LeftOrder.agda
new file mode 100644
index 0000000..5a31a5a
--- /dev/null
+++ b/src/Data/Sum/Relation/LeftOrder.agda
@@ -0,0 +1,220 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Sums of binary relations
+------------------------------------------------------------------------
+
+module Data.Sum.Relation.LeftOrder where
+
+open import Data.Sum as Sum
+import Data.Sum.Relation.Core as Core
+open import Data.Sum.Relation.Pointwise as Pointwise using (Pointwise; ₁≁₂)
+open import Data.Product
+open import Data.Unit.Base using (⊤)
+open import Data.Empty
+open import Function
+open import Level
+open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+
+module _ {a₁ a₂} {A₁ : Set a₁} {A₂ : Set a₂} where
+
+ ----------------------------------------------------------------------
+ -- Left order
+
+ open Core public using (₁∼₂; ₁∼₁; ₂∼₂)
+
+ infixr 1 _⊎-<_
+
+ _⊎-<_ : ∀ {ℓ₁ ℓ₂} → Rel A₁ ℓ₁ → Rel A₂ ℓ₂ → Rel (A₁ ⊎ A₂) _
+ _⊎-<_ = Core.⊎ʳ ⊤
+
+ ----------------------------------------------------------------------
+ -- Some properties which are preserved by _⊎-<_
+
+ module _ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} where
+
+ ⊎-<-refl : Reflexive ∼₁ → Reflexive ∼₂ →
+ Reflexive (∼₁ ⊎-< ∼₂)
+ ⊎-<-refl refl₁ refl₂ = Core._⊎-refl_ refl₁ refl₂
+
+ ⊎-<-transitive : Transitive ∼₁ → Transitive ∼₂ →
+ Transitive (∼₁ ⊎-< ∼₂)
+ ⊎-<-transitive trans₁ trans₂ = Core._⊎-transitive_ trans₁ trans₂
+
+ ⊎-<-asymmetric : Asymmetric ∼₁ → Asymmetric ∼₂ →
+ Asymmetric (∼₁ ⊎-< ∼₂)
+ ⊎-<-asymmetric asym₁ asym₂ = asym₁ Core.⊎-asymmetric asym₂
+
+ ⊎-<-total : Total ∼₁ → Total ∼₂ → Total (∼₁ ⊎-< ∼₂)
+ ⊎-<-total total₁ total₂ = total
+ where
+ total : Total (_ ⊎-< _)
+ total (inj₁ x) (inj₁ y) = Sum.map ₁∼₁ ₁∼₁ $ total₁ x y
+ total (inj₂ x) (inj₂ y) = Sum.map ₂∼₂ ₂∼₂ $ total₂ x y
+ total (inj₁ x) (inj₂ y) = inj₁ (₁∼₂ _)
+ total (inj₂ x) (inj₁ y) = inj₂ (₁∼₂ _)
+
+ ⊎-<-decidable : Decidable ∼₁ → Decidable ∼₂ →
+ (∀ {x y} → Dec (inj₁ x ⟨ ∼₁ ⊎-< ∼₂ ⟩ inj₂ y)) →
+ Decidable (∼₁ ⊎-< ∼₂)
+ ⊎-<-decidable dec₁ dec₂ dec₁₂ = Core.⊎-decidable dec₁ dec₂ dec₁₂
+
+
+ module _ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {≈₁ : Rel A₁ ℓ₂}
+ {ℓ₃ ℓ₄} {∼₂ : Rel A₂ ℓ₃} {≈₂ : Rel A₂ ℓ₄} where
+
+ ⊎-<-reflexive : ≈₁ ⇒ ∼₁ → ≈₂ ⇒ ∼₂ →
+ (Pointwise ≈₁ ≈₂) ⇒ (∼₁ ⊎-< ∼₂)
+ ⊎-<-reflexive refl₁ refl₂ = refl₁ Core.⊎-reflexive refl₂
+
+ ⊎-<-irreflexive : Irreflexive ≈₁ ∼₁ → Irreflexive ≈₂ ∼₂ →
+ Irreflexive (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-irreflexive irrefl₁ irrefl₂ = irrefl₁ Core.⊎-irreflexive irrefl₂
+
+ ⊎-<-antisymmetric : Antisymmetric ≈₁ ∼₁ → Antisymmetric ≈₂ ∼₂ →
+ Antisymmetric (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-antisymmetric antisym₁ antisym₂ = antisym₁ Core.⊎-antisymmetric antisym₂
+
+ ⊎-<-respects₂ : ∼₁ Respects₂ ≈₁ → ∼₂ Respects₂ ≈₂ →
+ (∼₁ ⊎-< ∼₂) Respects₂ (Pointwise ≈₁ ≈₂)
+ ⊎-<-respects₂ resp₁ resp₂ = Core._⊎-≈-respects₂_ resp₁ resp₂
+
+ ⊎-<-trichotomous : Trichotomous ≈₁ ∼₁ → Trichotomous ≈₂ ∼₂ →
+ Trichotomous (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-trichotomous tri₁ tri₂ = tri
+ where
+ tri : Trichotomous (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ tri (inj₁ x) (inj₂ y) = tri< (₁∼₂ _) ₁≁₂ (λ())
+ tri (inj₂ x) (inj₁ y) = tri> (λ()) (λ()) (₁∼₂ _)
+ tri (inj₁ x) (inj₁ y) with tri₁ x y
+ ... | tri< x<y x≉y x≯y =
+ tri< (₁∼₁ x<y) (x≉y ∘ Core.drop-inj₁) (x≯y ∘ Core.drop-inj₁)
+ ... | tri≈ x≮y x≈y x≯y =
+ tri≈ (x≮y ∘ Core.drop-inj₁) (₁∼₁ x≈y) (x≯y ∘ Core.drop-inj₁)
+ ... | tri> x≮y x≉y x>y =
+ tri> (x≮y ∘ Core.drop-inj₁) (x≉y ∘ Core.drop-inj₁) (₁∼₁ x>y)
+ tri (inj₂ x) (inj₂ y) with tri₂ x y
+ ... | tri< x<y x≉y x≯y =
+ tri< (₂∼₂ x<y) (x≉y ∘ Core.drop-inj₂) (x≯y ∘ Core.drop-inj₂)
+ ... | tri≈ x≮y x≈y x≯y =
+ tri≈ (x≮y ∘ Core.drop-inj₂) (₂∼₂ x≈y) (x≯y ∘ Core.drop-inj₂)
+ ... | tri> x≮y x≉y x>y =
+ tri> (x≮y ∘ Core.drop-inj₂) (x≉y ∘ Core.drop-inj₂) (₂∼₂ x>y)
+
+ ----------------------------------------------------------------------
+ -- Some collections of properties which are preserved
+
+ module _ {ℓ₁ ℓ₂} {≈₁ : Rel A₁ ℓ₁} {∼₁ : Rel A₁ ℓ₂}
+ {ℓ₃ ℓ₄} {≈₂ : Rel A₂ ℓ₃} {∼₂ : Rel A₂ ℓ₄} where
+
+ ⊎-<-isPreorder : IsPreorder ≈₁ ∼₁ → IsPreorder ≈₂ ∼₂ →
+ IsPreorder (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-isPreorder pre₁ pre₂ = record
+ { isEquivalence =
+ Pointwise.⊎-isEquivalence (isEquivalence pre₁) (isEquivalence pre₂)
+ ; reflexive = ⊎-<-reflexive (reflexive pre₁) (reflexive pre₂)
+ ; trans = ⊎-<-transitive (trans pre₁) (trans pre₂)
+ }
+ where open IsPreorder
+
+ ⊎-<-isPartialOrder : IsPartialOrder ≈₁ ∼₁ →
+ IsPartialOrder ≈₂ ∼₂ →
+ IsPartialOrder (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-isPartialOrder po₁ po₂ = record
+ { isPreorder = ⊎-<-isPreorder (isPreorder po₁) (isPreorder po₂)
+ ; antisym = ⊎-<-antisymmetric (antisym po₁) (antisym po₂)
+ }
+ where open IsPartialOrder
+
+ ⊎-<-isStrictPartialOrder : IsStrictPartialOrder ≈₁ ∼₁ →
+ IsStrictPartialOrder ≈₂ ∼₂ →
+ IsStrictPartialOrder (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-isStrictPartialOrder spo₁ spo₂ = record
+ { isEquivalence = Pointwise.⊎-isEquivalence (isEquivalence spo₁) (isEquivalence spo₂)
+ ; irrefl = ⊎-<-irreflexive (irrefl spo₁) (irrefl spo₂)
+ ; trans = ⊎-<-transitive (trans spo₁) (trans spo₂)
+ ; <-resp-≈ = ⊎-<-respects₂ (<-resp-≈ spo₁) (<-resp-≈ spo₂)
+ }
+ where open IsStrictPartialOrder
+
+ ⊎-<-isTotalOrder : IsTotalOrder ≈₁ ∼₁ →
+ IsTotalOrder ≈₂ ∼₂ →
+ IsTotalOrder (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-isTotalOrder to₁ to₂ = record
+ { isPartialOrder = ⊎-<-isPartialOrder (isPartialOrder to₁) (isPartialOrder to₂)
+ ; total = ⊎-<-total (total to₁) (total to₂)
+ }
+ where open IsTotalOrder
+
+ ⊎-<-isDecTotalOrder : IsDecTotalOrder ≈₁ ∼₁ →
+ IsDecTotalOrder ≈₂ ∼₂ →
+ IsDecTotalOrder (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-isDecTotalOrder to₁ to₂ = record
+ { isTotalOrder = ⊎-<-isTotalOrder (isTotalOrder to₁) (isTotalOrder to₂)
+ ; _≟_ = Pointwise.⊎-decidable (_≟_ to₁) (_≟_ to₂)
+ ; _≤?_ = ⊎-<-decidable (_≤?_ to₁) (_≤?_ to₂) (yes (₁∼₂ _))
+ }
+ where open IsDecTotalOrder
+
+ ⊎-<-isStrictTotalOrder : IsStrictTotalOrder ≈₁ ∼₁ →
+ IsStrictTotalOrder ≈₂ ∼₂ →
+ IsStrictTotalOrder (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
+ ⊎-<-isStrictTotalOrder sto₁ sto₂ = record
+ { isEquivalence = Pointwise.⊎-isEquivalence (isEquivalence sto₁) (isEquivalence sto₂)
+ ; trans = ⊎-<-transitive (trans sto₁) (trans sto₂)
+ ; compare = ⊎-<-trichotomous (compare sto₁) (compare sto₂)
+ }
+ where open IsStrictTotalOrder
+
+------------------------------------------------------------------------
+-- "Packages" can also be combined.
+
+module _ {a b c d e f} where
+
+ ⊎-<-preorder : Preorder a b c →
+ Preorder d e f →
+ Preorder _ _ _
+ ⊎-<-preorder p₁ p₂ = record
+ { isPreorder =
+ ⊎-<-isPreorder (isPreorder p₁) (isPreorder p₂)
+ } where open Preorder
+
+ ⊎-<-poset : Poset a b c →
+ Poset a b c →
+ Poset _ _ _
+ ⊎-<-poset po₁ po₂ = record
+ { isPartialOrder =
+ ⊎-<-isPartialOrder (isPartialOrder po₁) (isPartialOrder po₂)
+ } where open Poset
+
+ ⊎-<-strictPartialOrder : StrictPartialOrder a b c →
+ StrictPartialOrder d e f →
+ StrictPartialOrder _ _ _
+ ⊎-<-strictPartialOrder spo₁ spo₂ = record
+ { isStrictPartialOrder =
+ ⊎-<-isStrictPartialOrder (isStrictPartialOrder spo₁) (isStrictPartialOrder spo₂)
+ } where open StrictPartialOrder
+
+ ⊎-<-totalOrder : TotalOrder a b c →
+ TotalOrder d e f →
+ TotalOrder _ _ _
+ ⊎-<-totalOrder to₁ to₂ = record
+ { isTotalOrder = ⊎-<-isTotalOrder (isTotalOrder to₁) (isTotalOrder to₂)
+ } where open TotalOrder
+
+ ⊎-<-decTotalOrder : DecTotalOrder a b c →
+ DecTotalOrder d e f →
+ DecTotalOrder _ _ _
+ ⊎-<-decTotalOrder to₁ to₂ = record
+ { isDecTotalOrder = ⊎-<-isDecTotalOrder (isDecTotalOrder to₁) (isDecTotalOrder to₂)
+ } where open DecTotalOrder
+
+ ⊎-<-strictTotalOrder : StrictTotalOrder a b c →
+ StrictTotalOrder a b c →
+ StrictTotalOrder _ _ _
+ ⊎-<-strictTotalOrder sto₁ sto₂ = record
+ { isStrictTotalOrder = ⊎-<-isStrictTotalOrder (isStrictTotalOrder sto₁) (isStrictTotalOrder sto₂)
+ } where open StrictTotalOrder
diff --git a/src/Data/Sum/Relation/Pointwise.agda b/src/Data/Sum/Relation/Pointwise.agda
new file mode 100644
index 0000000..1bd5e62
--- /dev/null
+++ b/src/Data/Sum/Relation/Pointwise.agda
@@ -0,0 +1,389 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Pointwise sum
+------------------------------------------------------------------------
+
+module Data.Sum.Relation.Pointwise where
+
+open import Data.Sum as Sum
+import Data.Sum.Relation.Core as Core
+open import Data.Empty using (⊥)
+open import Function
+open import Function.Equality as F
+ using (_⟶_; _⟨$⟩_)
+open import Function.Equivalence as Eq
+ using (Equivalence; _⇔_; module Equivalence)
+open import Function.Injection as Inj
+ using (Injection; _↣_; module Injection)
+open import Function.Inverse as Inv
+ using (Inverse; _↔_; module Inverse)
+open import Function.LeftInverse as LeftInv
+ using (LeftInverse; _↞_; module LeftInverse)
+open import Function.Related
+open import Function.Surjection as Surj
+ using (Surjection; _↠_; module Surjection)
+open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+
+module _ {a₁ a₂} {A₁ : Set a₁} {A₂ : Set a₂} where
+
+ ----------------------------------------------------------------------
+ -- Pointwise sum
+
+ open Core public using (₁∼₂; ₁∼₁; ₂∼₂)
+
+ Pointwise : ∀ {ℓ₁ ℓ₂} → Rel A₁ ℓ₁ → Rel A₂ ℓ₂ → Rel (A₁ ⊎ A₂) _
+ Pointwise = Core.⊎ʳ ⊥
+
+ ----------------------------------------------------------------------
+ -- Helpers
+
+ ₁≁₂ : ∀ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} →
+ ∀ {x y} → ¬ (Pointwise ∼₁ ∼₂ (inj₁ x) (inj₂ y))
+ ₁≁₂ (₁∼₂ ())
+
+ ----------------------------------------------------------------------
+ -- Some properties which are preserved by Pointwise
+
+ module _ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {∼₂ : Rel A₂ ℓ₂} where
+
+ ⊎-refl : Reflexive ∼₁ → Reflexive ∼₂ →
+ Reflexive (Pointwise ∼₁ ∼₂)
+ ⊎-refl refl₁ refl₂ = Core._⊎-refl_ refl₁ refl₂
+
+ ⊎-symmetric : Symmetric ∼₁ → Symmetric ∼₂ →
+ Symmetric (Pointwise ∼₁ ∼₂)
+ ⊎-symmetric sym₁ sym₂ = sym
+ where
+ sym : Symmetric (Pointwise _ _)
+ sym (₁∼₁ x₁∼y₁) = ₁∼₁ (sym₁ x₁∼y₁)
+ sym (₂∼₂ x₂∼y₂) = ₂∼₂ (sym₂ x₂∼y₂)
+ sym (₁∼₂ ())
+
+ ⊎-transitive : Transitive ∼₁ → Transitive ∼₂ →
+ Transitive (Pointwise ∼₁ ∼₂)
+ ⊎-transitive trans₁ trans₂ = Core._⊎-transitive_ trans₁ trans₂
+
+ ⊎-asymmetric : Asymmetric ∼₁ → Asymmetric ∼₂ →
+ Asymmetric (Pointwise ∼₁ ∼₂)
+ ⊎-asymmetric asym₁ asym₂ = Core._⊎-asymmetric_ asym₁ asym₂
+
+ ⊎-substitutive : ∀ {ℓ₃} → Substitutive ∼₁ ℓ₃ → Substitutive ∼₂ ℓ₃ →
+ Substitutive (Pointwise ∼₁ ∼₂) ℓ₃
+ ⊎-substitutive subst₁ subst₂ = subst
+ where
+ subst : Substitutive (Pointwise _ _) _
+ subst P (₁∼₁ x∼y) Px = subst₁ (λ z → P (inj₁ z)) x∼y Px
+ subst P (₂∼₂ x∼y) Px = subst₂ (λ z → P (inj₂ z)) x∼y Px
+ subst P (₁∼₂ ()) Px
+
+ ⊎-decidable : Decidable ∼₁ → Decidable ∼₂ →
+ Decidable (Pointwise ∼₁ ∼₂)
+ ⊎-decidable dec₁ dec₂ = Core.⊎-decidable dec₁ dec₂ (no ₁≁₂)
+
+ module _ {ℓ₁ ℓ₂} {∼₁ : Rel A₁ ℓ₁} {≈₁ : Rel A₁ ℓ₂}
+ {ℓ₃ ℓ₄} {∼₂ : Rel A₂ ℓ₃} {≈₂ : Rel A₂ ℓ₄} where
+
+ ⊎-reflexive : ≈₁ ⇒ ∼₁ → ≈₂ ⇒ ∼₂ →
+ (Pointwise ≈₁ ≈₂) ⇒ (Pointwise ∼₁ ∼₂)
+ ⊎-reflexive refl₁ refl₂ = Core._⊎-reflexive_ refl₁ refl₂
+
+ ⊎-irreflexive : Irreflexive ≈₁ ∼₁ → Irreflexive ≈₂ ∼₂ →
+ Irreflexive (Pointwise ≈₁ ≈₂) (Pointwise ∼₁ ∼₂)
+ ⊎-irreflexive irrefl₁ irrefl₂ =
+ Core._⊎-irreflexive_ irrefl₁ irrefl₂
+
+ ⊎-antisymmetric : Antisymmetric ≈₁ ∼₁ → Antisymmetric ≈₂ ∼₂ →
+ Antisymmetric (Pointwise ≈₁ ≈₂) (Pointwise ∼₁ ∼₂)
+ ⊎-antisymmetric antisym₁ antisym₂ =
+ Core._⊎-antisymmetric_ antisym₁ antisym₂
+
+ ⊎-respects₂ : ∼₁ Respects₂ ≈₁ → ∼₂ Respects₂ ≈₂ →
+ (Pointwise ∼₁ ∼₂) Respects₂ (Pointwise ≈₁ ≈₂)
+ ⊎-respects₂ resp₁ resp₂ = Core._⊎-≈-respects₂_ resp₁ resp₂
+
+ ----------------------------------------------------------------------
+ -- Some collections of properties which are preserved
+
+ module _ {ℓ₁ ℓ₂} {≈₁ : Rel A₁ ℓ₁} {≈₂ : Rel A₂ ℓ₂} where
+
+ ⊎-isEquivalence : IsEquivalence ≈₁ → IsEquivalence ≈₂ →
+ IsEquivalence (Pointwise ≈₁ ≈₂)
+ ⊎-isEquivalence eq₁ eq₂ = record
+ { refl = ⊎-refl (refl eq₁) (refl eq₂)
+ ; sym = ⊎-symmetric (sym eq₁) (sym eq₂)
+ ; trans = ⊎-transitive (trans eq₁) (trans eq₂)
+ }
+ where open IsEquivalence
+
+ ⊎-isDecEquivalence : IsDecEquivalence ≈₁ → IsDecEquivalence ≈₂ →
+ IsDecEquivalence (Pointwise ≈₁ ≈₂)
+ ⊎-isDecEquivalence eq₁ eq₂ = record
+ { isEquivalence =
+ ⊎-isEquivalence (isEquivalence eq₁) (isEquivalence eq₂)
+ ; _≟_ = ⊎-decidable (_≟_ eq₁) (_≟_ eq₂)
+ }
+ where open IsDecEquivalence
+
+ module _ {ℓ₁ ℓ₂} {≈₁ : Rel A₁ ℓ₁} {∼₁ : Rel A₁ ℓ₂}
+ {ℓ₃ ℓ₄} {≈₂ : Rel A₂ ℓ₃} {∼₂ : Rel A₂ ℓ₄} where
+
+ ⊎-isPreorder : IsPreorder ≈₁ ∼₁ → IsPreorder ≈₂ ∼₂ →
+ IsPreorder (Pointwise ≈₁ ≈₂) (Pointwise ∼₁ ∼₂)
+ ⊎-isPreorder pre₁ pre₂ = record
+ { isEquivalence =
+ ⊎-isEquivalence (isEquivalence pre₁) (isEquivalence pre₂)
+ ; reflexive = ⊎-reflexive (reflexive pre₁) (reflexive pre₂)
+ ; trans = ⊎-transitive (trans pre₁) (trans pre₂)
+ }
+ where open IsPreorder
+
+ ⊎-isPartialOrder : IsPartialOrder ≈₁ ∼₁ →
+ IsPartialOrder ≈₂ ∼₂ →
+ IsPartialOrder
+ (Pointwise ≈₁ ≈₂) (Pointwise ∼₁ ∼₂)
+ ⊎-isPartialOrder po₁ po₂ = record
+ { isPreorder = ⊎-isPreorder (isPreorder po₁) (isPreorder po₂)
+ ; antisym = ⊎-antisymmetric (antisym po₁) (antisym po₂)
+ }
+ where open IsPartialOrder
+
+ ⊎-isStrictPartialOrder : IsStrictPartialOrder ≈₁ ∼₁ →
+ IsStrictPartialOrder ≈₂ ∼₂ →
+ IsStrictPartialOrder
+ (Pointwise ≈₁ ≈₂) (Pointwise ∼₁ ∼₂)
+ ⊎-isStrictPartialOrder spo₁ spo₂ = record
+ { isEquivalence =
+ ⊎-isEquivalence (isEquivalence spo₁) (isEquivalence spo₂)
+ ; irrefl = ⊎-irreflexive (irrefl spo₁) (irrefl spo₂)
+ ; trans = ⊎-transitive (trans spo₁) (trans spo₂)
+ ; <-resp-≈ = ⊎-respects₂ (<-resp-≈ spo₁) (<-resp-≈ spo₂)
+ }
+ where open IsStrictPartialOrder
+
+------------------------------------------------------------------------
+-- "Packages" can also be combined.
+
+module _ {a b c d} where
+
+ ⊎-setoid : Setoid a b → Setoid c d → Setoid _ _
+ ⊎-setoid s₁ s₂ = record
+ { isEquivalence =
+ ⊎-isEquivalence (isEquivalence s₁) (isEquivalence s₂)
+ } where open Setoid
+
+ ⊎-decSetoid : DecSetoid a b → DecSetoid c d → DecSetoid _ _
+ ⊎-decSetoid ds₁ ds₂ = record
+ { isDecEquivalence =
+ ⊎-isDecEquivalence (isDecEquivalence ds₁) (isDecEquivalence ds₂)
+ } where open DecSetoid
+
+ -- Some additional notation for combining setoids
+ infix 4 _⊎ₛ_
+ _⊎ₛ_ : Setoid a b → Setoid c d → Setoid _ _
+ _⊎ₛ_ = ⊎-setoid
+
+module _ {a b c d e f} where
+
+ ⊎-preorder : Preorder a b c → Preorder d e f → Preorder _ _ _
+ ⊎-preorder p₁ p₂ = record
+ { isPreorder =
+ ⊎-isPreorder (isPreorder p₁) (isPreorder p₂)
+ } where open Preorder
+
+ ⊎-poset : Poset a b c → Poset a b c → Poset _ _ _
+ ⊎-poset po₁ po₂ = record
+ { isPartialOrder =
+ ⊎-isPartialOrder (isPartialOrder po₁) (isPartialOrder po₂)
+ } where open Poset
+
+------------------------------------------------------------------------
+-- The propositional equality setoid over products can be
+-- decomposed using ×-Rel
+
+Pointwise-≡⇒≡ : ∀ {a b} {A : Set a} {B : Set b} →
+ (Pointwise _≡_ _≡_) ⇒ _≡_ {A = A ⊎ B}
+Pointwise-≡⇒≡ (₁∼₂ ())
+Pointwise-≡⇒≡ (₁∼₁ P.refl) = P.refl
+Pointwise-≡⇒≡ (₂∼₂ P.refl) = P.refl
+
+≡⇒Pointwise-≡ : ∀ {a b} {A : Set a} {B : Set b} →
+ _≡_ {A = A ⊎ B} ⇒ (Pointwise _≡_ _≡_)
+≡⇒Pointwise-≡ P.refl = ⊎-refl P.refl P.refl
+
+Pointwise-≡↔≡ : ∀ {a b} (A : Set a) (B : Set b) →
+ Inverse ((P.setoid A) ⊎ₛ (P.setoid B))
+ (P.setoid (A ⊎ B))
+Pointwise-≡↔≡ _ _ = record
+ { to = record { _⟨$⟩_ = id; cong = Pointwise-≡⇒≡ }
+ ; from = record { _⟨$⟩_ = id; cong = ≡⇒Pointwise-≡ }
+ ; inverse-of = record
+ { left-inverse-of = λ _ → ⊎-refl P.refl P.refl
+ ; right-inverse-of = λ _ → P.refl
+ }
+ }
+
+_⊎-≟_ : ∀ {a b} {A : Set a} {B : Set b} →
+ Decidable {A = A} _≡_ → Decidable {A = B} _≡_ →
+ Decidable {A = A ⊎ B} _≡_
+(dec₁ ⊎-≟ dec₂) s₁ s₂ = Dec.map′ Pointwise-≡⇒≡ ≡⇒Pointwise-≡ (s₁ ≟ s₂)
+ where
+ open DecSetoid (⊎-decSetoid (P.decSetoid dec₁) (P.decSetoid dec₂))
+
+------------------------------------------------------------------------
+-- Setoid "relatedness"
+
+module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
+ where
+
+ _⊎-⟶_ : A ⟶ B → C ⟶ D → (A ⊎ₛ C) ⟶ (B ⊎ₛ D)
+ _⊎-⟶_ f g = record
+ { _⟨$⟩_ = fg
+ ; cong = fg-cong
+ }
+ where
+ open Setoid (A ⊎ₛ C) using () renaming (_≈_ to _≈AC_)
+ open Setoid (B ⊎ₛ D) using () renaming (_≈_ to _≈BD_)
+
+ fg = Sum.map (_⟨$⟩_ f) (_⟨$⟩_ g)
+
+ fg-cong : _≈AC_ =[ fg ]⇒ _≈BD_
+ fg-cong (₁∼₂ ())
+ fg-cong (₁∼₁ x∼₁y) = ₁∼₁ $ F.cong f x∼₁y
+ fg-cong (₂∼₂ x∼₂y) = ₂∼₂ $ F.cong g x∼₂y
+
+module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
+ where
+
+ _⊎-equivalence_ : Equivalence A B → Equivalence C D →
+ Equivalence (A ⊎ₛ C) (B ⊎ₛ D)
+ A⇔B ⊎-equivalence C⇔D = record
+ { to = to A⇔B ⊎-⟶ to C⇔D
+ ; from = from A⇔B ⊎-⟶ from C⇔D
+ } where open Equivalence
+
+ _⊎-injection_ : Injection A B → Injection C D →
+ Injection (A ⊎ₛ C) (B ⊎ₛ D)
+ _⊎-injection_ A↣B C↣D = record
+ { to = to A↣B ⊎-⟶ to C↣D
+ ; injective = inj _ _
+ }
+ where
+ open Injection
+ open Setoid (A ⊎ₛ C) using () renaming (_≈_ to _≈AC_)
+ open Setoid (B ⊎ₛ D) using () renaming (_≈_ to _≈BD_)
+
+ inj : ∀ x y →
+ (to A↣B ⊎-⟶ to C↣D) ⟨$⟩ x ≈BD (to A↣B ⊎-⟶ to C↣D) ⟨$⟩ y →
+ x ≈AC y
+ inj (inj₁ x) (inj₁ y) (₁∼₁ x∼₁y) = ₁∼₁ (injective A↣B x∼₁y)
+ inj (inj₂ x) (inj₂ y) (₂∼₂ x∼₂y) = ₂∼₂ (injective C↣D x∼₂y)
+ inj (inj₁ x) (inj₂ y) (₁∼₂ ())
+ inj (inj₂ x) (inj₁ y) ()
+
+ _⊎-left-inverse_ : LeftInverse A B → LeftInverse C D →
+ LeftInverse (A ⊎ₛ C) (B ⊎ₛ D)
+ A↞B ⊎-left-inverse C↞D = record
+ { to = Equivalence.to eq
+ ; from = Equivalence.from eq
+ ; left-inverse-of = [ ₁∼₁ ∘ left-inverse-of A↞B
+ , ₂∼₂ ∘ left-inverse-of C↞D
+ ]
+ }
+ where
+ open LeftInverse
+ eq = LeftInverse.equivalence A↞B ⊎-equivalence
+ LeftInverse.equivalence C↞D
+
+module _ {a₁ a₂ b₁ b₂ c₁ c₂ d₁ d₂}
+ {A : Setoid a₁ a₂} {B : Setoid b₁ b₂}
+ {C : Setoid c₁ c₂} {D : Setoid d₁ d₂}
+ where
+
+ _⊎-surjection_ : Surjection A B → Surjection C D →
+ Surjection (A ⊎ₛ C) (B ⊎ₛ D)
+ A↠B ⊎-surjection C↠D = record
+ { to = LeftInverse.from inv
+ ; surjective = record
+ { from = LeftInverse.to inv
+ ; right-inverse-of = LeftInverse.left-inverse-of inv
+ }
+ }
+ where
+ open Surjection
+ inv = right-inverse A↠B ⊎-left-inverse right-inverse C↠D
+
+ _⊎-inverse_ : Inverse A B → Inverse C D →
+ Inverse (A ⊎ₛ C) (B ⊎ₛ D)
+ A↔B ⊎-inverse C↔D = record
+ { to = Surjection.to surj
+ ; from = Surjection.from surj
+ ; inverse-of = record
+ { left-inverse-of = LeftInverse.left-inverse-of inv
+ ; right-inverse-of = Surjection.right-inverse-of surj
+ }
+ }
+ where
+ open Inverse
+ surj = Inverse.surjection A↔B ⊎-surjection
+ Inverse.surjection C↔D
+ inv = Inverse.left-inverse A↔B ⊎-left-inverse
+ Inverse.left-inverse C↔D
+
+------------------------------------------------------------------------
+-- Propositional "relatedness"
+
+module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
+
+ _⊎-⇔_ : A ⇔ B → C ⇔ D → (A ⊎ C) ⇔ (B ⊎ D)
+ _⊎-⇔_ A⇔B C⇔D =
+ Inverse.equivalence (Pointwise-≡↔≡ B D) ⟨∘⟩
+ (A⇔B ⊎-equivalence C⇔D) ⟨∘⟩
+ Eq.sym (Inverse.equivalence (Pointwise-≡↔≡ A C))
+ where open Eq using () renaming (_∘_ to _⟨∘⟩_)
+
+ _⊎-↣_ : A ↣ B → C ↣ D → (A ⊎ C) ↣ (B ⊎ D)
+ _⊎-↣_ A↣B C↣D =
+ Inverse.injection (Pointwise-≡↔≡ B D) ⟨∘⟩
+ (A↣B ⊎-injection C↣D) ⟨∘⟩
+ Inverse.injection (Inv.sym (Pointwise-≡↔≡ A C))
+ where open Inj using () renaming (_∘_ to _⟨∘⟩_)
+
+ _⊎-↞_ : A ↞ B → C ↞ D → (A ⊎ C) ↞ (B ⊎ D)
+ _⊎-↞_ A↞B C↞D =
+ Inverse.left-inverse (Pointwise-≡↔≡ B D) ⟨∘⟩
+ (A↞B ⊎-left-inverse C↞D) ⟨∘⟩
+ Inverse.left-inverse (Inv.sym (Pointwise-≡↔≡ A C))
+ where open LeftInv using () renaming (_∘_ to _⟨∘⟩_)
+
+ _⊎-↠_ : A ↠ B → C ↠ D → (A ⊎ C) ↠ (B ⊎ D)
+ _⊎-↠_ A↠B C↠D =
+ Inverse.surjection (Pointwise-≡↔≡ B D) ⟨∘⟩
+ (A↠B ⊎-surjection C↠D) ⟨∘⟩
+ Inverse.surjection (Inv.sym (Pointwise-≡↔≡ A C))
+ where open Surj using () renaming (_∘_ to _⟨∘⟩_)
+
+ _⊎-↔_ : A ↔ B → C ↔ D → (A ⊎ C) ↔ (B ⊎ D)
+ _⊎-↔_ A↔B C↔D =
+ Pointwise-≡↔≡ B D ⟨∘⟩
+ (A↔B ⊎-inverse C↔D) ⟨∘⟩
+ Inv.sym (Pointwise-≡↔≡ A C)
+ where open Inv using () renaming (_∘_ to _⟨∘⟩_)
+
+_⊎-cong_ : ∀ {k a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
+ A ∼[ k ] B → C ∼[ k ] D → (A ⊎ C) ∼[ k ] (B ⊎ D)
+_⊎-cong_ {implication} = Sum.map
+_⊎-cong_ {reverse-implication} = λ f g → lam (Sum.map (app-← f) (app-← g))
+_⊎-cong_ {equivalence} = _⊎-⇔_
+_⊎-cong_ {injection} = _⊎-↣_
+_⊎-cong_ {reverse-injection} = λ f g → lam (app-↢ f ⊎-↣ app-↢ g)
+_⊎-cong_ {left-inverse} = _⊎-↞_
+_⊎-cong_ {surjection} = _⊎-↠_
+_⊎-cong_ {bijection} = _⊎-↔_
diff --git a/src/Data/Table.agda b/src/Data/Table.agda
new file mode 100644
index 0000000..bc15e37
--- /dev/null
+++ b/src/Data/Table.agda
@@ -0,0 +1,36 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Fixed-size tables of values, implemented as functions from 'Fin n'.
+-- Similar to 'Data.Vec', but focusing on ease of retrieval instead of
+-- ease of adding and removing elements.
+------------------------------------------------------------------------
+
+module Data.Table where
+
+open import Data.Table.Base public
+
+open import Data.Bool using (true; false)
+open import Data.Fin using (Fin; _≟_)
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Inverse using (Inverse; _↔_)
+open import Relation.Nullary using (yes; no)
+open import Relation.Nullary.Decidable using (⌊_⌋)
+
+--------------------------------------------------------------------------------
+-- Combinators
+--------------------------------------------------------------------------------
+
+-- Changes the order of elements in the table according to a permutation (i.e.
+-- an 'Inverse' object on the indices).
+
+permute : ∀ {m n a} {A : Set a} → Fin m ↔ Fin n → Table A n → Table A m
+permute π = rearrange (Inverse.to π ⟨$⟩_)
+
+-- The result of 'select z i t' takes the value of 'lookup t i' at index 'i',
+-- and 'z' everywhere else.
+
+select : ∀ {n} {a} {A : Set a} → A → Fin n → Table A n → Table A n
+lookup (select z i t) j with j ≟ i
+... | yes _ = lookup t i
+... | no _ = z
diff --git a/src/Data/Table/Base.agda b/src/Data/Table/Base.agda
new file mode 100644
index 0000000..3c71186
--- /dev/null
+++ b/src/Data/Table/Base.agda
@@ -0,0 +1,93 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Tables, basic types and operations
+------------------------------------------------------------------------
+
+module Data.Table.Base where
+
+open import Data.Nat
+open import Data.Fin
+open import Data.Product using (_×_ ; _,_)
+open import Data.List as List using (List)
+open import Data.Vec as Vec using (Vec)
+open import Function using (_∘_; flip)
+
+------------------------------------------------------------------------
+-- Type
+
+record Table {a} (A : Set a) n : Set a where
+ constructor tabulate
+ field lookup : Fin n → A
+open Table public
+
+------------------------------------------------------------------------
+-- Basic operations
+
+module _ {a} {A : Set a} where
+
+ head : ∀ {n} → Table A (suc n) → A
+ head t = lookup t zero
+
+ tail : ∀ {n} → Table A (suc n) → Table A n
+ tail t = tabulate (lookup t ∘ suc)
+
+ uncons : ∀ {n} → Table A (suc n) → A × Table A n
+ uncons t = head t , tail t
+
+ remove : ∀ {n} → Fin (suc n) → Table A (suc n) → Table A n
+ remove i t = tabulate (lookup t ∘ punchIn i)
+
+------------------------------------------------------------------------
+-- Operations for transforming tables
+
+module _ {a} {A : Set a} where
+
+ rearrange : ∀ {m n} → (Fin m → Fin n) → Table A n → Table A m
+ rearrange f t = tabulate (lookup t ∘ f)
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ map : ∀ {n} → (A → B) → Table A n → Table B n
+ map f t = tabulate (f ∘ lookup t)
+
+ _⊛_ : ∀ {n} → Table (A → B) n → Table A n → Table B n
+ fs ⊛ xs = tabulate λ i → lookup fs i (lookup xs i)
+
+------------------------------------------------------------------------
+-- Operations for reducing tables
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ foldr : ∀ {n} → (A → B → B) → B → Table A n → B
+ foldr {n = zero} f z t = z
+ foldr {n = suc n} f z t = f (head t) (foldr f z (tail t))
+
+ foldl : ∀ {n} → (B → A → B) → B → Table A n → B
+ foldl {n = zero} f z t = z
+ foldl {n = suc n} f z t = foldl f (f z (head t)) (tail t)
+
+------------------------------------------------------------------------
+-- Operations for building tables
+
+module _ {a} {A : Set a} where
+
+ replicate : ∀ {n} → A → Table A n
+ replicate x = tabulate (λ _ → x)
+
+------------------------------------------------------------------------
+-- Operations for converting tables
+
+module _ {a} {A : Set a} where
+
+ toList : ∀ {n} → Table A n → List A
+ toList = List.tabulate ∘ lookup
+
+ fromList : ∀ (xs : List A) → Table A (List.length xs)
+ fromList = tabulate ∘ List.lookup
+
+ fromVec : ∀ {n} → Vec A n → Table A n
+ fromVec = tabulate ∘ flip Vec.lookup
+
+ toVec : ∀ {n} → Table A n → Vec A n
+ toVec = Vec.tabulate ∘ lookup
diff --git a/src/Data/Table/Properties.agda b/src/Data/Table/Properties.agda
new file mode 100644
index 0000000..f21e7aa
--- /dev/null
+++ b/src/Data/Table/Properties.agda
@@ -0,0 +1,117 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Table-related properties
+------------------------------------------------------------------------
+
+module Data.Table.Properties where
+
+open import Data.Table
+open import Data.Table.Relation.Equality
+
+open import Data.Bool using (true; false; if_then_else_)
+open import Data.Nat using (zero; suc)
+open import Data.Empty using (⊥-elim)
+open import Data.Fin using (Fin; suc; zero; _≟_; punchIn)
+import Data.Fin.Properties as FP
+open import Data.Fin.Permutation as Perm using (Permutation; _⟨$⟩ʳ_; _⟨$⟩ˡ_)
+open import Data.List as L using (List; _∷_; [])
+open import Data.List.Any using (here; there; index)
+open import Data.List.Membership.Propositional using (_∈_)
+open import Data.Product as Product using (Σ; ∃; _,_; proj₁; proj₂)
+open import Data.Vec as V using (Vec; _∷_; [])
+import Data.Vec.Properties as VP
+open import Function using (_∘_; flip)
+open import Function.Inverse using (Inverse)
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≢_; refl; sym; cong)
+open import Relation.Nullary using (yes; no)
+open import Relation.Nullary.Negation using (contradiction)
+
+
+------------------------------------------------------------------------
+-- select
+
+module _ {a} {A : Set a} where
+
+ -- Selecting from any table is the same as selecting from a constant table.
+
+ select-const : ∀ {n} (z : A) (i : Fin n) t →
+ select z i t ≗ select z i (replicate (lookup t i))
+ select-const z i t j with j ≟ i
+ ... | yes _ = refl
+ ... | no _ = refl
+
+ -- Selecting an element from a table then looking it up is the same as looking
+ -- up the index in the original table
+
+ select-lookup : ∀ {n x i} (t : Table A n) →
+ lookup (select x i t) i ≡ lookup t i
+ select-lookup {i = i} t with i ≟ i
+ ... | yes _ = refl
+ ... | no i≢i = contradiction refl i≢i
+
+ -- Selecting an element from a table then removing the same element produces a
+ -- constant table
+
+ select-remove : ∀ {n x} i (t : Table A (suc n)) →
+ remove i (select x i t) ≗ replicate {n = n} x
+ select-remove i t j with punchIn i j ≟ i
+ ... | yes p = contradiction p (FP.punchInᵢ≢i _ _)
+ ... | no ¬p = refl
+
+
+------------------------------------------------------------------------
+-- permute
+
+ -- Removing an index 'i' from a table permuted with 'π' is the same as
+ -- removing the element, then permuting with 'π' minus 'i'.
+
+ remove-permute : ∀ {m n} (π : Permutation (suc m) (suc n))
+ i (t : Table A (suc n)) →
+ remove (π ⟨$⟩ˡ i) (permute π t)
+ ≗ permute (Perm.remove (π ⟨$⟩ˡ i) π) (remove i t)
+ remove-permute π i t j = P.cong (lookup t) (Perm.punchIn-permute′ π i j)
+
+------------------------------------------------------------------------
+-- fromList
+
+module _ {a} {A : Set a} where
+
+ fromList-∈ : ∀ {xs : List A} (i : Fin (L.length xs)) → lookup (fromList xs) i ∈ xs
+ fromList-∈ {[]} ()
+ fromList-∈ {x ∷ xs} zero = here refl
+ fromList-∈ {x ∷ xs} (suc i) = there (fromList-∈ i)
+
+ index-fromList-∈ : ∀ {xs i} → index (fromList-∈ {xs} i) ≡ i
+ index-fromList-∈ {[]} {()}
+ index-fromList-∈ {x ∷ xs} {zero} = refl
+ index-fromList-∈ {x ∷ xs} {suc i} = cong suc index-fromList-∈
+
+ fromList-index : ∀ {xs} {x : A} (x∈xs : x ∈ xs) → lookup (fromList xs) (index x∈xs) ≡ x
+ fromList-index (here px) = sym px
+ fromList-index (there x∈xs) = fromList-index x∈xs
+
+
+------------------------------------------------------------------------
+-- There exists an isomorphism between tables and vectors.
+
+module _ {a n} {A : Set a} where
+
+ ↔Vec : Inverse (≡-setoid A n) (P.setoid (Vec A n))
+ ↔Vec = record
+ { to = record { _⟨$⟩_ = toVec ; cong = VP.tabulate-cong }
+ ; from = P.→-to-⟶ fromVec
+ ; inverse-of = record
+ { left-inverse-of = VP.lookup∘tabulate ∘ lookup
+ ; right-inverse-of = VP.tabulate∘lookup
+ }
+ }
+
+------------------------------------------------------------------------
+-- Other
+
+module _ {a} {A : Set a} where
+
+ lookup∈ : ∀ {xs : List A} (i : Fin (L.length xs)) → ∃ λ x → x ∈ xs
+ lookup∈ i = _ , fromList-∈ i
diff --git a/src/Data/Table/Relation/Equality.agda b/src/Data/Table/Relation/Equality.agda
new file mode 100644
index 0000000..d71cdea
--- /dev/null
+++ b/src/Data/Table/Relation/Equality.agda
@@ -0,0 +1,33 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Pointwise table equality
+------------------------------------------------------------------------
+
+module Data.Table.Relation.Equality where
+
+open import Relation.Binary using (Setoid)
+open import Data.Table.Base
+open import Data.Nat using (ℕ)
+open import Function using (_∘_)
+open import Relation.Binary.PropositionalEquality
+ as P using (_≡_; _→-setoid_)
+
+setoid : ∀ {c p} → Setoid c p → ℕ → Setoid _ _
+setoid S n = record
+ { Carrier = Table Carrier n
+ ; _≈_ = λ t t′ → ∀ i → lookup t i ≈ lookup t′ i
+ ; isEquivalence = record
+ { refl = λ i → refl
+ ; sym = λ p → sym ∘ p
+ ; trans = λ p q i → trans (p i) (q i)
+ }
+ }
+ where open Setoid S
+
+≡-setoid : ∀ {a} → Set a → ℕ → Setoid _ _
+≡-setoid A = setoid (P.setoid A)
+
+module _ {a} {A : Set a} {n} where
+ open Setoid (≡-setoid A n) public
+ using () renaming (_≈_ to _≗_)
diff --git a/src/Data/These.agda b/src/Data/These.agda
new file mode 100644
index 0000000..a4d6104
--- /dev/null
+++ b/src/Data/These.agda
@@ -0,0 +1,77 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- An either-or-both data type
+------------------------------------------------------------------------
+
+module Data.These where
+
+open import Level
+open import Function
+
+data These {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
+ this : A → These A B
+ that : B → These A B
+ these : A → B → These A B
+
+-- map
+
+map : ∀ {a₁ a₂ b₁ b₂} {A₁ : Set a₁} {A₂ : Set a₂} {B₁ : Set b₁} {B₂ : Set b₂}
+ (f : A₁ → A₂) (g : B₁ → B₂) → These A₁ B₁ → These A₂ B₂
+map f g (this a) = this (f a)
+map f g (that b) = that (g b)
+map f g (these a b) = these (f a) (g b)
+
+map₁ : ∀ {a₁ a₂ b} {A₁ : Set a₁} {A₂ : Set a₂} {B : Set b}
+ (f : A₁ → A₂) → These A₁ B → These A₂ B
+map₁ f = map f id
+
+map₂ : ∀ {a b₁ b₂} {A : Set a} {B₁ : Set b₁} {B₂ : Set b₂}
+ (g : B₁ → B₂) → These A B₁ → These A B₂
+map₂ = map id
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+-- fold
+
+ fold : ∀ {c} {C : Set c} → (A → C) → (B → C) → (A → B → C) → These A B → C
+ fold l r lr (this a) = l a
+ fold l r lr (that b) = r b
+ fold l r lr (these a b) = lr a b
+
+-- swap
+
+ swap : These A B → These B A
+ swap = fold that this (flip these)
+
+-- align
+
+module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
+
+ alignWith : ∀ {e f} {E : Set e} {F : Set f} →
+ (These A C → E) → (These B D → F) → These A B → These C D → These E F
+ alignWith f g (this a) (this c) = this (f (these a c))
+ alignWith f g (this a) (that d) = these (f (this a)) (g (that d))
+ alignWith f g (this a) (these c d) = these (f (these a c)) (g (that d))
+ alignWith f g (that b) (this c) = these (f (that c)) (g (this b))
+ alignWith f g (that b) (that d) = that (g (these b d))
+ alignWith f g (that b) (these c d) = these (f (that c)) (g (these b d))
+ alignWith f g (these a b) (this c) = these (f (these a c)) (g (this b))
+ alignWith f g (these a b) (that d) = these (f (this a)) (g (these b d))
+ alignWith f g (these a b) (these c d) = these (f (these a c)) (g (these b d))
+
+ align : These A B → These C D → These (These A C) (These B D)
+ align = alignWith id id
+
+-- projections
+
+module _ {a} {A : Set a} where
+
+ leftMost : These A A → A
+ leftMost = fold id id const
+
+ rightMost : These A A → A
+ rightMost = fold id id (flip const)
+
+ mergeThese : (A → A → A) → These A A → A
+ mergeThese = fold id id
diff --git a/src/Data/These/Categorical/Left.agda b/src/Data/These/Categorical/Left.agda
new file mode 100644
index 0000000..764f08f
--- /dev/null
+++ b/src/Data/These/Categorical/Left.agda
@@ -0,0 +1,54 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Left-biased universe-sensitive functor and monad instances for These.
+--
+-- To minimize the universe level of the RawFunctor, we require that
+-- elements of B are "lifted" to a copy of B at a higher universe level
+-- (a ⊔ b).
+-- See the Data.Product.Categorical.Examples for how this is done in a
+-- Product-based similar setting.
+------------------------------------------------------------------------
+
+-- This functor can be understood as a notion of computation which can
+-- either fail (this), succeed (that) or accumulate warnings whilst
+-- delivering a successful computation (these).
+
+-- It is a good alternative to Data.Product.Categorical when the notion
+-- of warnings does not have a neutral element (e.g. List⁺).
+
+open import Level
+open import Algebra
+
+module Data.These.Categorical.Left {c ℓ} (W : Semigroup c ℓ) (b : Level) where
+
+open Semigroup W
+open import Data.These.Categorical.Left.Base Carrier b public
+
+open import Data.These
+open import Category.Applicative
+open import Category.Monad
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+applicative : RawApplicative Theseₗ
+applicative = record
+ { pure = that
+ ; _⊛_ = ap
+ } where
+
+ ap : ∀ {A B}→ Theseₗ (A → B) → Theseₗ A → Theseₗ B
+ ap (this w) t = this w
+ ap (that f) t = map₂ f t
+ ap (these w f) t = map (w ∙_) f t
+
+monad : RawMonad Theseₗ
+monad = record
+ { return = that
+ ; _>>=_ = bind
+ } where
+
+ bind : ∀ {A B} → Theseₗ A → (A → Theseₗ B) → Theseₗ B
+ bind (this w) f = this w
+ bind (that t) f = f t
+ bind (these w t) f = map₁ (w ∙_) (f t)
diff --git a/src/Data/These/Categorical/Left/Base.agda b/src/Data/These/Categorical/Left/Base.agda
new file mode 100644
index 0000000..4c4f0ba
--- /dev/null
+++ b/src/Data/These/Categorical/Left/Base.agda
@@ -0,0 +1,59 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Base definitions for the left-biased universe-sensitive functor and
+-- monad instances for These.
+--
+-- To minimize the universe level of the RawFunctor, we require that
+-- elements of B are "lifted" to a copy of B at a higher universe level
+-- (a ⊔ b).
+-- See the Data.Product.Categorical.Examples for how this is done in a
+-- Product-based similar setting.
+------------------------------------------------------------------------
+
+open import Level
+
+module Data.These.Categorical.Left.Base {a} (A : Set a) (b : Level) where
+
+open import Data.These
+open import Category.Functor
+open import Category.Applicative
+open import Category.Monad
+open import Function
+
+Theseₗ : Set (a ⊔ b) → Set (a ⊔ b)
+Theseₗ B = These A B
+
+functor : RawFunctor Theseₗ
+functor = record { _<$>_ = map₂ }
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {F} (App : RawApplicative {a ⊔ b} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → Theseₗ (F A) → F (Theseₗ A)
+ sequenceA (this a) = pure (this a)
+ sequenceA (that b) = that <$> b
+ sequenceA (these a b) = these a <$> b
+
+ mapA : ∀ {A B} → (A → F B) → Theseₗ A → F (Theseₗ B)
+ mapA f = sequenceA ∘ map₂ f
+
+ forA : ∀ {A B} → Theseₗ A → (A → F B) → F (Theseₗ B)
+ forA = flip mapA
+
+module _ {M} (Mon : RawMonad {a ⊔ b} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM : ∀ {A} → Theseₗ (M A) → M (Theseₗ A)
+ sequenceM = sequenceA App
+
+ mapM : ∀ {A B} → (A → M B) → Theseₗ A → M (Theseₗ B)
+ mapM = mapA App
+
+ forM : ∀ {A B} → Theseₗ A → (A → M B) → M (Theseₗ B)
+ forM = forA App
diff --git a/src/Data/These/Categorical/Right.agda b/src/Data/These/Categorical/Right.agda
new file mode 100644
index 0000000..7c622e3
--- /dev/null
+++ b/src/Data/These/Categorical/Right.agda
@@ -0,0 +1,54 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Right-biased universe-sensitive functor and monad instances for These.
+--
+-- To minimize the universe level of the RawFunctor, we require that
+-- elements of B are "lifted" to a copy of B at a higher universe level
+-- (a ⊔ b).
+-- See the Data.Product.Categorical.Examples for how this is done in a
+-- Product-based similar setting.
+------------------------------------------------------------------------
+
+-- This functor can be understood as a notion of computation which can
+-- either fail (that), succeed (this) or accumulate warnings whilst
+-- delivering a successful computation (these).
+
+-- It is a good alternative to Data.Product.Categorical when the notion
+-- of warnings does not have a neutral element (e.g. List⁺).
+
+open import Level
+open import Algebra
+
+module Data.These.Categorical.Right (a : Level) {c ℓ} (W : Semigroup c ℓ) where
+
+open Semigroup W
+open import Data.These.Categorical.Right.Base a Carrier public
+
+open import Data.These
+open import Category.Applicative
+open import Category.Monad
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+applicative : RawApplicative Theseᵣ
+applicative = record
+ { pure = this
+ ; _⊛_ = ap
+ } where
+
+ ap : ∀ {A B}→ Theseᵣ (A → B) → Theseᵣ A → Theseᵣ B
+ ap (this f) t = map₁ f t
+ ap (that w) t = that w
+ ap (these f w) t = map f (w ∙_) t
+
+monad : RawMonad Theseᵣ
+monad = record
+ { return = this
+ ; _>>=_ = bind
+ } where
+
+ bind : ∀ {A B} → Theseᵣ A → (A → Theseᵣ B) → Theseᵣ B
+ bind (this t) f = f t
+ bind (that w) f = that w
+ bind (these t w) f = map₂ (w ∙_) (f t)
diff --git a/src/Data/These/Categorical/Right/Base.agda b/src/Data/These/Categorical/Right/Base.agda
new file mode 100644
index 0000000..84bac16
--- /dev/null
+++ b/src/Data/These/Categorical/Right/Base.agda
@@ -0,0 +1,59 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Base definitions for the right-biased universe-sensitive functor and
+-- monad instances for These.
+--
+-- To minimize the universe level of the RawFunctor, we require that
+-- elements of B are "lifted" to a copy of B at a higher universe level
+-- (a ⊔ b).
+-- See the Data.Product.Categorical.Examples for how this is done in a
+-- Product-based similar setting.
+------------------------------------------------------------------------
+
+open import Level
+
+module Data.These.Categorical.Right.Base (a : Level) {b} (B : Set b) where
+
+open import Data.These
+open import Category.Functor
+open import Category.Applicative
+open import Category.Monad
+open import Function
+
+Theseᵣ : Set (a ⊔ b) → Set (a ⊔ b)
+Theseᵣ A = These A B
+
+functor : RawFunctor Theseᵣ
+functor = record { _<$>_ = map₁ }
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {F} (App : RawApplicative {a ⊔ b} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A} → Theseᵣ (F A) → F (Theseᵣ A)
+ sequenceA (this a) = this <$> a
+ sequenceA (that b) = pure (that b)
+ sequenceA (these a b) = flip these b <$> a
+
+ mapA : ∀ {A B} → (A → F B) → Theseᵣ A → F (Theseᵣ B)
+ mapA f = sequenceA ∘ map₁ f
+
+ forA : ∀ {A B} → Theseᵣ A → (A → F B) → F (Theseᵣ B)
+ forA = flip mapA
+
+module _ {M} (Mon : RawMonad {a ⊔ b} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM : ∀ {A} → Theseᵣ (M A) → M (Theseᵣ A)
+ sequenceM = sequenceA App
+
+ mapM : ∀ {A B} → (A → M B) → Theseᵣ A → M (Theseᵣ B)
+ mapM = mapA App
+
+ forM : ∀ {A B} → Theseᵣ A → (A → M B) → M (Theseᵣ B)
+ forM = forA App
diff --git a/src/Data/Vec.agda b/src/Data/Vec.agda
index 7be843a..10befc6 100644
--- a/src/Data/Vec.agda
+++ b/src/Data/Vec.agda
@@ -6,14 +6,15 @@
module Data.Vec where
-open import Category.Functor
-open import Category.Applicative
open import Data.Nat
open import Data.Fin using (Fin; zero; suc)
open import Data.List.Base as List using (List)
open import Data.Product as Prod using (∃; ∃₂; _×_; _,_)
+open import Data.These as These using (These; this; that; these)
open import Function
open import Relation.Binary.PropositionalEquality using (_≡_; refl)
+open import Relation.Nullary using (yes; no)
+open import Relation.Unary using (Pred; Decidable)
------------------------------------------------------------------------
-- Types
@@ -24,12 +25,6 @@ data Vec {a} (A : Set a) : ℕ → Set a where
[] : Vec A zero
_∷_ : ∀ {n} (x : A) (xs : Vec A n) → Vec A (suc n)
-infix 4 _∈_
-
-data _∈_ {a} {A : Set a} : A → {n : ℕ} → Vec A n → Set a where
- here : ∀ {n} {x} {xs : Vec A n} → x ∈ x ∷ xs
- there : ∀ {n} {x y} {xs : Vec A n} (x∈xs : x ∈ xs) → x ∈ y ∷ xs
-
infix 4 _[_]=_
data _[_]=_ {a} {A : Set a} :
@@ -39,7 +34,7 @@ data _[_]=_ {a} {A : Set a} :
(xs[i]=x : xs [ i ]= x) → y ∷ xs [ suc i ]= x
------------------------------------------------------------------------
--- Some operations
+-- Basic operations
head : ∀ {a n} {A : Set a} → Vec A (1 + n) → A
head (x ∷ xs) = x
@@ -47,8 +42,37 @@ head (x ∷ xs) = x
tail : ∀ {a n} {A : Set a} → Vec A (1 + n) → Vec A n
tail (x ∷ xs) = xs
-[_] : ∀ {a} {A : Set a} → A → Vec A 1
-[ x ] = x ∷ []
+lookup : ∀ {a n} {A : Set a} → Fin n → Vec A n → A
+lookup zero (x ∷ xs) = x
+lookup (suc i) (x ∷ xs) = lookup i xs
+
+insert : ∀ {a n} {A : Set a} → Fin (suc n) → A → Vec A n → Vec A (suc n)
+insert zero x xs = x ∷ xs
+insert (suc ()) x []
+insert (suc i) x (y ∷ xs) = y ∷ insert i x xs
+
+remove : ∀ {a n} {A : Set a} → Fin (suc n) → Vec A (suc n) → Vec A n
+remove zero (_ ∷ xs) = xs
+remove (suc ()) (x ∷ [])
+remove (suc i) (x ∷ y ∷ xs) = x ∷ remove i (y ∷ xs)
+
+-- Update.
+
+infixl 6 _[_]≔_
+
+_[_]≔_ : ∀ {a n} {A : Set a} → Vec A n → Fin n → A → Vec A n
+(x ∷ xs) [ zero ]≔ y = y ∷ xs
+(x ∷ xs) [ suc i ]≔ y = x ∷ xs [ i ]≔ y
+
+------------------------------------------------------------------------
+-- Operations for transforming vectors
+
+map : ∀ {a b n} {A : Set a} {B : Set b} →
+ (A → B) → Vec A n → Vec B n
+map f [] = []
+map f (x ∷ xs) = f x ∷ map f xs
+
+-- Concatenation.
infixr 5 _++_
@@ -56,6 +80,50 @@ _++_ : ∀ {a m n} {A : Set a} → Vec A m → Vec A n → Vec A (m + n)
[] ++ ys = ys
(x ∷ xs) ++ ys = x ∷ (xs ++ ys)
+concat : ∀ {a m n} {A : Set a} →
+ Vec (Vec A m) n → Vec A (n * m)
+concat [] = []
+concat (xs ∷ xss) = xs ++ concat xss
+
+-- Align and Zip.
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
+
+ alignWith : ∀ {m n} → (These A B → C) → Vec A m → Vec B n → Vec C (m ⊔ n)
+ alignWith f [] bs = map (f ∘′ that) bs
+ alignWith f as@(_ ∷ _) [] = map (f ∘′ this) as
+ alignWith f (a ∷ as) (b ∷ bs) = f (these a b) ∷ alignWith f as bs
+
+ zipWith : ∀ {n} → (A → B → C) → Vec A n → Vec B n → Vec C n
+ zipWith f [] [] = []
+ zipWith f (x ∷ xs) (y ∷ ys) = f x y ∷ zipWith f xs ys
+
+ unzipWith : ∀ {n} → (A → B × C) → Vec A n → Vec B n × Vec C n
+ unzipWith f [] = [] , []
+ unzipWith f (a ∷ as) = Prod.zip _∷_ _∷_ (f a) (unzipWith f as)
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ align : ∀ {m n} → Vec A m → Vec B n → Vec (These A B) (m ⊔ n)
+ align = alignWith id
+
+ zip : ∀ {n} → Vec A n → Vec B n → Vec (A × B) n
+ zip = zipWith _,_
+
+ unzip : ∀ {n} → Vec (A × B) n → Vec A n × Vec B n
+ unzip = unzipWith id
+
+-- Interleaving.
+
+infixr 5 _⋎_
+
+_⋎_ : ∀ {a m n} {A : Set a} →
+ Vec A m → Vec A n → Vec A (m +⋎ n)
+[] ⋎ ys = ys
+(x ∷ xs) ⋎ ys = x ∷ (ys ⋎ xs)
+
+-- Pointwise application
+
infixl 4 _⊛_
_⊛_ : ∀ {a b n} {A : Set a} {B : Set b} →
@@ -63,39 +131,26 @@ _⊛_ : ∀ {a b n} {A : Set a} {B : Set b} →
[] ⊛ _ = []
(f ∷ fs) ⊛ (x ∷ xs) = f x ∷ (fs ⊛ xs)
-replicate : ∀ {a n} {A : Set a} → A → Vec A n
-replicate {n = zero} x = []
-replicate {n = suc n} x = x ∷ replicate x
+-- Multiplication
-applicative : ∀ {a n} → RawApplicative (λ (A : Set a) → Vec A n)
-applicative = record
- { pure = replicate
- ; _⊛_ = _⊛_
- }
+infixl 1 _>>=_
-map : ∀ {a b n} {A : Set a} {B : Set b} →
- (A → B) → Vec A n → Vec B n
-map f [] = []
-map f (x ∷ xs) = f x ∷ map f xs
+_>>=_ : ∀ {a b m n} {A : Set a} {B : Set b} →
+ Vec A m → (A → Vec B n) → Vec B (m * n)
+xs >>= f = concat (map f xs)
-functor : ∀ {a n} → RawFunctor (λ (A : Set a) → Vec A n)
-functor = record
- { _<$>_ = map
- }
+infixl 4 _⊛*_
-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 f [] [] = []
-zipWith f (x ∷ xs) (y ∷ ys) = f x y ∷ zipWith f xs ys
+_⊛*_ : ∀ {a b m n} {A : Set a} {B : Set b} →
+ Vec (A → B) m → Vec A n → Vec B (m * n)
+fs ⊛* xs = fs >>= λ f → map f xs
-zip : ∀ {a b n} {A : Set a} {B : Set b} →
- Vec A n → Vec B n → Vec (A × B) n
-zip = zipWith _,_
+allPairs : ∀ {a b m n} {A : Set a} {B : Set b} →
+ Vec A m → Vec B n → Vec (A × B) (m * n)
+allPairs xs ys = map _,_ xs ⊛* ys
-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)
+------------------------------------------------------------------------
+-- Operations for reducing vectors
foldr : ∀ {a b} {A : Set a} (B : ℕ → Set b) {m} →
(∀ {n} → A → B n → B (suc n)) →
@@ -120,10 +175,37 @@ foldl₁ : ∀ {a} {A : Set a} {m} →
(A → A → A) → Vec A (suc m) → A
foldl₁ _⊕_ (x ∷ xs) = foldl _ _⊕_ x xs
-concat : ∀ {a m n} {A : Set a} →
- Vec (Vec A m) n → Vec A (n * m)
-concat [] = []
-concat (xs ∷ xss) = xs ++ concat xss
+-- Special folds
+
+sum : ∀ {n} → Vec ℕ n → ℕ
+sum = foldr _ _+_ 0
+
+count : ∀ {a p} {A : Set a} {P : Pred A p} → Decidable P →
+ ∀ {n} → Vec A n → ℕ
+count P? [] = zero
+count P? (x ∷ xs) with P? x
+... | yes _ = suc (count P? xs)
+... | no _ = count P? xs
+
+------------------------------------------------------------------------
+-- Operations for building vectors
+
+[_] : ∀ {a} {A : Set a} → A → Vec A 1
+[ x ] = x ∷ []
+
+replicate : ∀ {a n} {A : Set a} → A → Vec A n
+replicate {n = zero} x = []
+replicate {n = suc n} x = x ∷ replicate x
+
+tabulate : ∀ {n a} {A : Set a} → (Fin n → A) → Vec A n
+tabulate {zero} f = []
+tabulate {suc n} f = f zero ∷ tabulate (f ∘ suc)
+
+allFin : ∀ n → Vec (Fin n) n
+allFin _ = tabulate id
+
+------------------------------------------------------------------------
+-- Operations for dividing vectors
splitAt : ∀ {a} {A : Set a} m {n} (xs : Vec A (m + n)) →
∃₂ λ (ys : Vec A m) (zs : Vec A n) → xs ≡ ys ++ zs
@@ -148,18 +230,13 @@ group (suc n) k .(ys ++ zs) | (ys , zs , refl) with group n k zs
group (suc n) k .(ys ++ concat zss) | (ys , ._ , refl) | (zss , refl) =
((ys ∷ zss) , refl)
--- Splits a vector into two "halves".
-
split : ∀ {a n} {A : Set a} → Vec A n → Vec A ⌈ n /2⌉ × Vec A ⌊ n /2⌋
split [] = ([] , [])
split (x ∷ []) = (x ∷ [] , [])
split (x ∷ y ∷ xs) = Prod.map (_∷_ x) (_∷_ y) (split xs)
-reverse : ∀ {a n} {A : Set a} → Vec A n → Vec A n
-reverse {A = A} = foldl (Vec A) (λ rev x → x ∷ rev) []
-
-sum : ∀ {n} → Vec ℕ n → ℕ
-sum = foldr _ _+_ 0
+------------------------------------------------------------------------
+-- Operations for converting between lists
toList : ∀ {a n} {A : Set a} → Vec A n → List A
toList [] = List.[]
@@ -169,7 +246,11 @@ fromList : ∀ {a} {A : Set a} → (xs : List A) → Vec A (List.length xs)
fromList List.[] = []
fromList (List._∷_ x xs) = x ∷ fromList xs
--- Snoc.
+------------------------------------------------------------------------
+-- Operations for reversing vectors
+
+reverse : ∀ {a n} {A : Set a} → Vec A n → Vec A n
+reverse {A = A} = foldl (Vec A) (λ rev x → x ∷ rev) []
infixl 5 _∷ʳ_
@@ -191,59 +272,3 @@ init .(ys ∷ʳ y) | (ys , y , refl) = ys
last : ∀ {a n} {A : Set a} → Vec A (1 + n) → A
last xs with initLast xs
last .(ys ∷ʳ y) | (ys , y , refl) = y
-
--- Multiplying vectors
-
-infixl 1 _>>=_
-
-_>>=_ : ∀ {a b m n} {A : Set a} {B : Set b} →
- Vec A m → (A → Vec B n) → Vec B (m * n)
-xs >>= f = concat (map f xs)
-
-infixl 4 _⊛*_
-
-_⊛*_ : ∀ {a b m n} {A : Set a} {B : Set b} →
- Vec (A → B) m → Vec A n → Vec B (m * n)
-fs ⊛* xs = fs >>= λ f → map f xs
-
-allPairs : ∀ {a b} {A : Set a} {B : Set b} {m n}
- → Vec A m → Vec B n → Vec (A × B) (m * n)
-allPairs xs ys = map _,_ xs ⊛* ys
-
--- Interleaves the two vectors.
-
-infixr 5 _⋎_
-
-_⋎_ : ∀ {a m n} {A : Set a} →
- Vec A m → Vec A n → Vec A (m +⋎ n)
-[] ⋎ ys = ys
-(x ∷ xs) ⋎ ys = x ∷ (ys ⋎ xs)
-
--- A lookup function.
-
-lookup : ∀ {a n} {A : Set a} → Fin n → Vec A n → A
-lookup zero (x ∷ xs) = x
-lookup (suc i) (x ∷ xs) = lookup i xs
-
--- An inverse of flip lookup.
-
-tabulate : ∀ {n a} {A : Set a} → (Fin n → A) → Vec A n
-tabulate {zero} f = []
-tabulate {suc n} f = f zero ∷ tabulate (f ∘ suc)
-
--- Update.
-
-infixl 6 _[_]≔_
-
-_[_]≔_ : ∀ {a n} {A : Set a} → Vec A n → Fin n → A → Vec A n
-(x ∷ xs) [ zero ]≔ y = y ∷ xs
-(x ∷ xs) [ suc i ]≔ y = x ∷ xs [ i ]≔ y
-
--- Generates a vector containing all elements in Fin n. This function
--- is not placed in Data.Fin because Data.Vec depends on Data.Fin.
---
--- The implementation was suggested by Conor McBride ("Fwd: how to
--- count 0..n-1", http://thread.gmane.org/gmane.comp.lang.agda/2554).
-
-allFin : ∀ n → Vec (Fin n) n
-allFin _ = tabulate id
diff --git a/src/Data/Vec/All.agda b/src/Data/Vec/All.agda
index 0dfe1b1..4c7f23e 100644
--- a/src/Data/Vec/All.agda
+++ b/src/Data/Vec/All.agda
@@ -6,16 +6,16 @@
module Data.Vec.All where
-open import Data.Vec as Vec using (Vec; []; _∷_; zip)
-open import Data.Vec.Properties using (lookup-zip)
+open import Data.Nat using (zero; suc)
open import Data.Fin using (Fin; zero; suc)
+open import Data.Product as Prod using (_,_)
+open import Data.Vec as Vec using (Vec; []; _∷_)
open import Function using (_∘_)
open import Level using (_⊔_)
-open import Data.Product using (uncurry)
open import Relation.Nullary
import Relation.Nullary.Decidable as Dec
-open import Relation.Unary using (Decidable) renaming (_⊆_ to _⋐_)
-open import Relation.Binary.PropositionalEquality using (subst)
+open import Relation.Unary
+open import Relation.Binary.PropositionalEquality as P using (subst)
------------------------------------------------------------------------
-- All P xs means that all elements in xs satisfy P.
@@ -27,6 +27,9 @@ data All {a p} {A : Set a}
[] : All P []
_∷_ : ∀ {k x} {xs : Vec A k} (px : P x) (pxs : All P xs) → All P (x ∷ xs)
+------------------------------------------------------------------------
+-- Operations on All
+
head : ∀ {a p} {A : Set a} {P : A → Set p} {k x} {xs : Vec A k} →
All P (x ∷ xs) → P x
head (px ∷ pxs) = px
@@ -41,23 +44,16 @@ lookup () []
lookup zero (px ∷ pxs) = px
lookup (suc i) (px ∷ pxs) = lookup i pxs
-tabulate : ∀ {a p} {A : Set a} {P : A → Set p} {k} {xs : Vec A k} →
- (∀ x → P x) → All P xs
-tabulate {xs = []} hyp = []
-tabulate {xs = x ∷ xs} hyp = hyp x ∷ tabulate hyp
+tabulate : ∀ {a p} {A : Set a} {P : A → Set p} {k xs} →
+ (∀ i → P (Vec.lookup i xs)) → All P {k} xs
+tabulate {xs = []} pxs = []
+tabulate {xs = _ ∷ _} pxs = pxs zero ∷ tabulate (pxs ∘ suc)
map : ∀ {a p q} {A : Set a} {P : A → Set p} {Q : A → Set q} {k} →
- P ⋐ Q → All P {k} ⋐ All Q {k}
+ P ⊆ Q → All P {k} ⊆ All Q {k}
map g [] = []
map g (px ∷ pxs) = g px ∷ map g pxs
-all : ∀ {a p} {A : Set a} {P : A → Set p} {k} →
- Decidable P → Decidable (All P {k})
-all p [] = yes []
-all p (x ∷ xs) with p x
-all p (x ∷ xs) | yes px = Dec.map′ (_∷_ px) tail (all p xs)
-all p (x ∷ xs) | no ¬px = no (¬px ∘ head)
-
zipWith : ∀ {a b c p q r} {A : Set a} {B : Set b} {C : Set c} {_⊕_ : A → B → C}
{P : A → Set p} {Q : B → Set q} {R : C → Set r} →
(∀ {x y} → P x → Q y → R (x ⊕ y)) →
@@ -67,25 +63,36 @@ zipWith _⊕_ {xs = []} {[]} [] [] = []
zipWith _⊕_ {xs = x ∷ xs} {y ∷ ys} (px ∷ pxs) (qy ∷ qys) =
px ⊕ qy ∷ zipWith _⊕_ pxs qys
+zip : ∀ {a p q k} {A : Set a} {P : A → Set p} {Q : A → Set q} →
+ All P ∩ All Q ⊆ All (P ∩ Q) {k}
+zip ([] , []) = []
+zip (px ∷ pxs , qx ∷ qxs) = (px , qx) ∷ zip (pxs , qxs)
+
+unzip : ∀ {a p q k} {A : Set a} {P : A → Set p} {Q : A → Set q} →
+ All (P ∩ Q) {k} ⊆ All P ∩ All Q
+unzip [] = [] , []
+unzip (pqx ∷ pqxs) = Prod.zip _∷_ _∷_ pqx (unzip pqxs)
------------------------------------------------------------------------
--- All₂ P xs ys means that every pointwise pair in xs ys satisfy P.
-
-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 b p} {A : Set a} {B : Set b} {P : A → B → Set p} {k}
- {xs : Vec A k} {ys : Vec B k} →
- ∀ 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
+-- Properties of predicates preserved by All
+
+module _ {a p} {A : Set a} {P : A → Set p} where
+
+ all : ∀ {k} → Decidable P → Decidable (All P {k})
+ all P? [] = yes []
+ all P? (x ∷ xs) with P? x
+ ... | yes px = Dec.map′ (px ∷_) tail (all P? xs)
+ ... | no ¬px = no (¬px ∘ head)
+
+ universal : Universal P → ∀ {k} → Universal (All P {k})
+ universal u [] = []
+ universal u (x ∷ xs) = u x ∷ universal u xs
+
+ irrelevant : Irrelevant P → ∀ {k} → Irrelevant (All P {k})
+ irrelevant irr [] [] = P.refl
+ irrelevant irr (px₁ ∷ pxs₁) (px₂ ∷ pxs₂) =
+ P.cong₂ _∷_ (irr px₁ px₂) (irrelevant irr pxs₁ pxs₂)
+
+ satisfiable : Satisfiable P → ∀ {k} → Satisfiable (All P {k})
+ satisfiable (x , p) {zero} = [] , []
+ satisfiable (x , p) {suc k} = Prod.map (x ∷_) (p ∷_) (satisfiable (x , p))
diff --git a/src/Data/Vec/All/Properties.agda b/src/Data/Vec/All/Properties.agda
index 3af19c7..4b1f735 100644
--- a/src/Data/Vec/All/Properties.agda
+++ b/src/Data/Vec/All/Properties.agda
@@ -6,194 +6,175 @@
module Data.Vec.All.Properties where
-open import Data.Vec as Vec using (Vec; []; _∷_; zip; concat; _++_)
-open import Data.Vec.Properties using (map-id; lookup-zip)
+open import Data.List using ([]; _∷_)
+open import Data.List.All as List using ([]; _∷_)
open import Data.Product as Prod using (_×_; _,_; uncurry; uncurry′)
-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.
-
-All-map : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {k} {xs : Vec A k} →
- All (P ∘ f) xs → All P (Vec.map f xs)
-All-map [] = []
-All-map (px ∷ pxs) = px ∷ All-map pxs
-
-map-All : ∀ {a b p} {A : Set a} {B : Set b} {P : B → Set p}
- {f : A → B} {k} {xs : Vec A k} →
- All P (Vec.map f xs) → All (P ∘ f) xs
-map-All {xs = []} [] = []
-map-All {xs = _ ∷ _} (px ∷ pxs) = px ∷ map-All pxs
-
--- A variant of All.map.
-
-gmap : ∀ {a b p q}
- {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q}
- {f : A → B} {k} →
- P ⋐ Q ∘ f → All P {k} ⋐ All Q {k} ∘ Vec.map f
-gmap g = All-map ∘ All.map g
-
--- A variant of All-map for All₂.
-
-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 [] = []
-All₂-map (px ∷ pxs) = px ∷ All₂-map pxs
-
--- 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
+open import Data.Vec as Vec
+open import Data.Vec.All as All using (All; []; _∷_)
+open import Function using (_∘_; id)
+open import Function.Inverse using (_↔_; inverse)
+open import Relation.Unary using (Pred) renaming (_⊆_ to _⋐_)
+open import Relation.Binary.PropositionalEquality
+ using (_≡_; refl; cong; cong₂; →-to-⟶)
------------------------------------------------------------------------
--- Variants of gmap for All₂.
+-- map
-module _ {a b c p q} {A : Set a} {B : Set b} {C : Set c} where
+module _ {a b p} {A : Set a} {B : Set b} {P : Pred B p} {f : A → B} where
- -- A variant of gmap₂ shifting two functions from the binary
- -- relation to the vector.
+ map⁺ : ∀ {n} {xs : Vec A n} → All (P ∘ f) xs → All P (map f xs)
+ map⁺ [] = []
+ map⁺ (px ∷ pxs) = px ∷ map⁺ pxs
- 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
+ map⁻ : ∀ {n} {xs : Vec A n} → All P (map f xs) → All (P ∘ f) xs
+ map⁻ {xs = []} [] = []
+ map⁻ {xs = _ ∷ _} (px ∷ pxs) = px ∷ map⁻ pxs
- -- A variant of gmap₂ shifting only the first function from the binary
- -- relation to the vector.
+-- A variant of All.map
- 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
+module _ {a b p q} {A : Set a} {B : Set b} {f : A → B}
+ {P : Pred A p} {Q : Pred B q} where
+ gmap : ∀ {n} → P ⋐ Q ∘ f → All P {n} ⋐ All Q {n} ∘ map f
+ gmap g = map⁺ ∘ All.map g
------------------------------------------------------------------------
--- All and _++_
+-- _++_
-module _ {a n p} {A : Set a} {P : A → Set p} where
+module _ {a n p} {A : Set a} {P : Pred A p} where
- 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
+ ++⁺ : ∀ {m} {xs : Vec A m} {ys : Vec A n} →
+ 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 xs
- All-++ˡ⁻ [] _ = []
- All-++ˡ⁻ (x ∷ xs) (px ∷ pxs) = px ∷ All-++ˡ⁻ xs pxs
+ ++ˡ⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
+ All P (xs ++ ys) → All P xs
+ ++ˡ⁻ [] _ = []
+ ++ˡ⁻ (x ∷ xs) (px ∷ pxs) = px ∷ ++ˡ⁻ xs pxs
- 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
+ ++ʳ⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
+ All P (xs ++ ys) → All P ys
+ ++ʳ⁻ [] pys = pys
+ ++ʳ⁻ (x ∷ xs) (px ∷ pxs) = ++ʳ⁻ xs pxs
- All-++⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
+ ++⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
All P (xs ++ ys) → All P xs × All P ys
- All-++⁻ [] p = [] , p
- All-++⁻ (x ∷ xs) (px ∷ pxs) = Prod.map (px ∷_) id (All-++⁻ _ pxs)
+ ++⁻ [] p = [] , p
+ ++⁻ (x ∷ xs) (px ∷ pxs) = Prod.map₁ (px ∷_) (++⁻ _ 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)
+ ++⁺∘++⁻ : ∀ {m} (xs : Vec A m) {ys : Vec A n} →
+ (p : All P (xs ++ ys)) →
+ uncurry′ ++⁺ (++⁻ xs p) ≡ p
+ ++⁺∘++⁻ [] p = refl
+ ++⁺∘++⁻ (x ∷ xs) (px ∷ pxs) = cong (px ∷_) (++⁺∘++⁻ 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} →
+ (p : All P xs × All P ys) →
+ ++⁻ xs (uncurry ++⁺ p) ≡ p
+ ++⁻∘++⁺ ([] , pys) = refl
+ ++⁻∘++⁺ (px ∷ pxs , pys) rewrite ++⁻∘++⁺ (pxs , pys) = 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
- }
- }
+ ++↔ {xs = xs} = inverse (uncurry ++⁺) (++⁻ xs) ++⁻∘++⁺ (++⁺∘++⁻ xs)
------------------------------------------------------------------------
--- All₂ and _++_
-
-module _ {a b n p} {A : Set a} {B : Set b} {_~_ : REL A B p} where
+-- concat
- 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)
+module _ {a m p} {A : Set a} {P : Pred A p} where
- 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
+ concat⁺ : ∀ {n} {xss : Vec (Vec A m) n} →
+ All (All P) xss → All P (concat xss)
+ concat⁺ [] = []
+ concat⁺ (pxs ∷ pxss) = ++⁺ pxs (concat⁺ pxss)
- 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
+ concat⁻ : ∀ {n} (xss : Vec (Vec A m) n) →
+ All P (concat xss) → All (All P) xss
+ concat⁻ [] [] = []
+ concat⁻ (xs ∷ xss) pxss = ++ˡ⁻ xs pxss ∷ concat⁻ xss (++ʳ⁻ xs pxss)
------------------------------------------------------------------------
--- All and concat
+-- toList
-module _ {a m p} {A : Set a} {P : A → Set p} where
+module _ {a 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)
+ toList⁺ : ∀ {n} {xs : Vec A n} → List.All P (toList xs) → All P xs
+ toList⁺ {xs = []} [] = []
+ toList⁺ {xs = x ∷ xs} (px ∷ pxs) = px ∷ toList⁺ pxs
- 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)
+ toList⁻ : ∀ {n} {xs : Vec A n} → All P xs → List.All P (toList xs)
+ toList⁻ [] = []
+ toList⁻ (px ∷ pxs) = px ∷ toList⁻ pxs
------------------------------------------------------------------------
--- All₂ and concat
+-- fromList
-module _ {a b m p} {A : Set a} {B : Set b} {_~_ : REL A B p} where
+module _ {a p} {A : Set a} {P : A → Set 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)
+ fromList⁺ : ∀ {xs} → List.All P xs → All P (fromList xs)
+ fromList⁺ [] = []
+ fromList⁺ (px ∷ pxs) = px ∷ fromList⁺ pxs
- 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)
+ fromList⁻ : ∀ {xs} → All P (fromList xs) → List.All P xs
+ fromList⁻ {[]} [] = []
+ fromList⁻ {x ∷ xs} (px ∷ pxs) = px ∷ (fromList⁻ pxs)
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.16
+
+All-map = map⁺
+{-# WARNING_ON_USAGE All-map
+"Warning: All-map was deprecated in v0.16.
+Please use map⁺ instead."
+#-}
+map-All = map⁻
+{-# WARNING_ON_USAGE map-All
+"Warning: map-All was deprecated in v0.16.
+Please use map⁻ instead."
+#-}
+All-++⁺ = ++⁺
+{-# WARNING_ON_USAGE All-++⁺
+"Warning: All-++⁺ was deprecated in v0.16.
+Please use ++⁺ instead."
+#-}
+All-++ˡ⁻ = ++ˡ⁻
+{-# WARNING_ON_USAGE All-++ˡ⁻
+"Warning: All-++ˡ⁻ was deprecated in v0.16.
+Please use ++ˡ⁻ instead."
+#-}
+All-++ʳ⁻ = ++ʳ⁻
+{-# WARNING_ON_USAGE All-++ʳ⁻
+"Warning: All-++ʳ⁻ was deprecated in v0.16.
+Please use ++ʳ⁻ instead."
+#-}
+All-++⁻ = ++⁻
+{-# WARNING_ON_USAGE All-++⁻
+"Warning: All-++⁻ was deprecated in v0.16.
+Please use ++⁻ instead."
+#-}
+All-++⁺∘++⁻ = ++⁺∘++⁻
+{-# WARNING_ON_USAGE All-++⁺∘++⁻
+"Warning: All-++⁺∘++⁻ was deprecated in v0.16.
+Please use ++⁺∘++⁻ instead."
+#-}
+All-++⁻∘++⁺ = ++⁻∘++⁺
+{-# WARNING_ON_USAGE All-++⁻∘++⁺
+"Warning: All-++⁻∘++⁺ was deprecated in v0.16.
+Please use ++⁻∘++⁺ instead."
+#-}
+All-concat⁺ = concat⁺
+{-# WARNING_ON_USAGE All-concat⁺
+"Warning: All-concat⁺ was deprecated in v0.16.
+Please use concat⁺ instead."
+#-}
+All-concat⁻ = concat⁻
+{-# WARNING_ON_USAGE All-concat⁻
+"Warning: All-concat⁻ was deprecated in v0.16.
+Please use concat⁻ instead."
+#-}
diff --git a/src/Data/Vec/Any.agda b/src/Data/Vec/Any.agda
new file mode 100644
index 0000000..094bccf
--- /dev/null
+++ b/src/Data/Vec/Any.agda
@@ -0,0 +1,79 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Vectors where at least one element satisfies a given property
+------------------------------------------------------------------------
+
+module Data.Vec.Any {a} {A : Set a} where
+
+open import Data.Empty
+open import Data.Fin
+open import Data.Nat using (zero; suc)
+open import Data.Sum using (_⊎_; inj₁; inj₂; [_,_]′)
+open import Data.Vec as Vec using (Vec; []; [_]; _∷_)
+open import Data.Product as Prod using (∃; _,_)
+open import Level using (_⊔_)
+open import Relation.Nullary using (¬_; yes; no)
+open import Relation.Nullary.Negation using (contradiction)
+import Relation.Nullary.Decidable as Dec
+open import Relation.Unary
+
+------------------------------------------------------------------------
+-- Any P xs means that at least one element in xs satisfies P.
+
+data Any {p} (P : A → Set p) : ∀ {n} → Vec A n → Set (a ⊔ p) where
+ here : ∀ {n x} {xs : Vec A n} (px : P x) → Any P (x ∷ xs)
+ there : ∀ {n x} {xs : Vec A n} (pxs : Any P xs) → Any P (x ∷ xs)
+
+------------------------------------------------------------------------
+-- Operations on Any
+
+module _ {p} {P : A → Set p} {n x} {xs : Vec A n} where
+
+-- If the tail does not satisfy the predicate, then the head will.
+
+ head : ¬ Any P xs → Any P (x ∷ xs) → P x
+ head ¬pxs (here px) = px
+ head ¬pxs (there pxs) = contradiction pxs ¬pxs
+
+-- If the head does not satisfy the predicate, then the tail will.
+ tail : ¬ P x → Any P (x ∷ xs) → Any P xs
+ tail ¬px (here px) = ⊥-elim (¬px px)
+ tail ¬px (there pxs) = pxs
+
+-- Convert back and forth with sum
+ toSum : Any P (x ∷ xs) → P x ⊎ Any P xs
+ toSum (here px) = inj₁ px
+ toSum (there pxs) = inj₂ pxs
+
+ fromSum : P x ⊎ Any P xs → Any P (x ∷ xs)
+ fromSum = [ here , there ]′
+
+map : ∀ {p q} {P : A → Set p} {Q : A → Set q} →
+ P ⊆ Q → ∀ {n} → Any P {n} ⊆ Any Q {n}
+map g (here px) = here (g px)
+map g (there pxs) = there (map g pxs)
+
+index : ∀ {p} {P : A → Set p} {n} {xs : Vec A n} → Any P xs → Fin n
+index (here px) = zero
+index (there pxs) = suc (index pxs)
+
+-- If any element satisfies P, then P is satisfied.
+satisfied : ∀ {p} {P : A → Set p} {n} {xs : Vec A n} → Any P xs → ∃ P
+satisfied (here px) = _ , px
+satisfied (there pxs) = satisfied pxs
+
+------------------------------------------------------------------------
+-- Properties of predicates preserved by Any
+
+module _ {p} {P : A → Set p} where
+
+ any : Decidable P → ∀ {n} → Decidable (Any P {n})
+ any P? [] = no λ()
+ any P? (x ∷ xs) with P? x
+ ... | yes px = yes (here px)
+ ... | no ¬px = Dec.map′ there (tail ¬px) (any P? xs)
+
+ satisfiable : Satisfiable P → ∀ {n} → Satisfiable (Any P {suc n})
+ satisfiable (x , p) {zero} = x ∷ [] , here p
+ satisfiable (x , p) {suc n} = Prod.map (x ∷_) there (satisfiable (x , p))
diff --git a/src/Data/Vec/Categorical.agda b/src/Data/Vec/Categorical.agda
new file mode 100644
index 0000000..773d019
--- /dev/null
+++ b/src/Data/Vec/Categorical.agda
@@ -0,0 +1,81 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- A categorical view of Vec
+------------------------------------------------------------------------
+
+module Data.Vec.Categorical {a n} where
+
+open import Category.Applicative using (RawApplicative)
+open import Category.Applicative.Indexed using (Morphism)
+open import Category.Functor as Fun using (RawFunctor)
+import Function.Identity.Categorical as Id
+open import Category.Monad using (RawMonad)
+open import Data.Fin using (Fin)
+open import Data.Vec as Vec hiding (_⊛_)
+open import Data.Vec.Properties
+open import Function
+
+------------------------------------------------------------------------
+-- Functor and applicative
+
+functor : RawFunctor (λ (A : Set a) → Vec A n)
+functor = record
+ { _<$>_ = map
+ }
+
+applicative : RawApplicative (λ (A : Set a) → Vec A n)
+applicative = record
+ { pure = replicate
+ ; _⊛_ = Vec._⊛_
+ }
+
+------------------------------------------------------------------------
+-- Get access to other monadic functions
+
+module _ {f F} (App : RawApplicative {f} F) where
+
+ open RawApplicative App
+
+ sequenceA : ∀ {A n} → Vec (F A) n → F (Vec A n)
+ sequenceA [] = pure []
+ sequenceA (x ∷ xs) = _∷_ <$> x ⊛ sequenceA xs
+
+ mapA : ∀ {a} {A : Set a} {B n} → (A → F B) → Vec A n → F (Vec B n)
+ mapA f = sequenceA ∘ map f
+
+ forA : ∀ {a} {A : Set a} {B n} → Vec A n → (A → F B) → F (Vec B n)
+ forA = flip mapA
+
+module _ {m M} (Mon : RawMonad {m} M) where
+
+ private App = RawMonad.rawIApplicative Mon
+
+ sequenceM : ∀ {A n} → Vec (M A) n → M (Vec A n)
+ sequenceM = sequenceA App
+
+ mapM : ∀ {a} {A : Set a} {B n} → (A → M B) → Vec A n → M (Vec B n)
+ mapM = mapA App
+
+ forM : ∀ {a} {A : Set a} {B n} → Vec A n → (A → M B) → M (Vec B n)
+ forM = forA App
+
+------------------------------------------------------------------------
+-- Other
+
+-- lookup is a functor morphism from Vec to Identity.
+
+lookup-functor-morphism : (i : Fin n) → Fun.Morphism functor Id.functor
+lookup-functor-morphism i = record
+ { op = lookup i
+ ; op-<$> = lookup-map i
+ }
+
+-- lookup is an applicative functor morphism.
+
+lookup-morphism : (i : Fin n) → Morphism applicative Id.applicative
+lookup-morphism i = record
+ { op = lookup i
+ ; op-pure = lookup-replicate i
+ ; op-⊛ = lookup-⊛ i
+ }
diff --git a/src/Data/Vec/Equality.agda b/src/Data/Vec/Equality.agda
deleted file mode 100644
index 8933a6d..0000000
--- a/src/Data/Vec/Equality.agda
+++ /dev/null
@@ -1,105 +0,0 @@
-------------------------------------------------------------------------
--- The Agda standard library
---
--- Semi-heterogeneous vector equality
-------------------------------------------------------------------------
-
-module Data.Vec.Equality where
-
-open import Data.Vec
-open import Data.Nat.Base using (suc)
-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
- open module SS = Setoid S
- using () renaming (_≈_ to _≊_; Carrier to A)
-
- infix 4 _≈_
-
- data _≈_ : ∀ {n¹} → Vec A n¹ →
- ∀ {n²} → Vec A n² → Set (s₁ ⊔ s₂) where
- []-cong : [] ≈ []
- _∷-cong_ : ∀ {x¹ n¹} {xs¹ : Vec A n¹}
- {x² n²} {xs² : Vec A n²}
- (x¹≈x² : x¹ ≊ x²) (xs¹≈xs² : xs¹ ≈ xs²) →
- x¹ ∷ xs¹ ≈ x² ∷ xs²
-
- length-equal : ∀ {n¹} {xs¹ : Vec A n¹}
- {n²} {xs² : Vec A n²} →
- xs¹ ≈ xs² → n¹ ≡ n²
- length-equal []-cong = P.refl
- length-equal (_ ∷-cong eq₂) = P.cong suc $ length-equal eq₂
-
- refl : ∀ {n} (xs : Vec A n) → xs ≈ xs
- refl [] = []-cong
- refl (x ∷ xs) = SS.refl ∷-cong refl xs
-
- sym : ∀ {n m} {xs : Vec A n} {ys : Vec A m} →
- xs ≈ ys → ys ≈ xs
- sym []-cong = []-cong
- sym (x¹≡x² ∷-cong xs¹≈xs²) = SS.sym x¹≡x² ∷-cong sym xs¹≈xs²
-
- trans : ∀ {n m l} {xs : Vec A n} {ys : Vec A m} {zs : Vec A l} →
- xs ≈ ys → ys ≈ zs → xs ≈ zs
- trans []-cong []-cong = []-cong
- 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₂² →
- xs₁¹ ++ xs₂¹ ≈ xs₁² ++ xs₂²
- []-cong ++-cong eq₃ = eq₃
- (eq₁ ∷-cong eq₂) ++-cong eq₃ = eq₁ ∷-cong (eq₂ ++-cong eq₃)
-
-module DecidableEquality {d₁ d₂} (D : DecSetoid d₁ d₂) where
-
- private module DS = DecSetoid D
- open DS using () renaming (_≟_ to _≟′_ ; Carrier to A)
- open Equality DS.setoid
- open import Relation.Nullary
-
- infix 4 _≟_
-
- _≟_ : ∀ {n m} (xs : Vec A n) (ys : Vec A m) → Dec (xs ≈ ys)
- _≟_ [] [] = yes []-cong
- _≟_ [] (y ∷ ys) = no (λ())
- _≟_ (x ∷ xs) [] = no (λ())
- _≟_ (x ∷ xs) (y ∷ ys) with xs ≟ ys | x ≟′ y
- ... | yes xs≈ys | yes x≊y = yes (x≊y ∷-cong xs≈ys)
- ... | no ¬xs≈ys | _ = no helper
- where
- helper : ¬ (x ∷ xs ≈ y ∷ ys)
- helper (_ ∷-cong xs≈ys) = ¬xs≈ys xs≈ys
- ... | _ | no ¬x≊y = no helper
- where
- helper : ¬ (x ∷ xs ≈ y ∷ ys)
- helper (x≊y ∷-cong _) = ¬x≊y x≊y
-
-module PropositionalEquality {a} {A : Set a} where
-
- open Equality (P.setoid A) public
-
- to-≡ : ∀ {n} {xs ys : Vec A n} → xs ≈ ys → xs ≡ ys
- to-≡ []-cong = P.refl
- to-≡ (P.refl ∷-cong xs¹≈xs²) = P.cong (_∷_ _) $ to-≡ xs¹≈xs²
-
- from-≡ : ∀ {n} {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/Membership/Propositional.agda b/src/Data/Vec/Membership/Propositional.agda
new file mode 100644
index 0000000..01251f1
--- /dev/null
+++ b/src/Data/Vec/Membership/Propositional.agda
@@ -0,0 +1,24 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Membership of vectors based on propositional equality,
+-- along with some additional definitions.
+------------------------------------------------------------------------
+
+module Data.Vec.Membership.Propositional {a} {A : Set a} where
+
+open import Data.Vec using (Vec)
+open import Data.Vec.Any using (Any)
+open import Relation.Binary.PropositionalEquality using (_≡_)
+open import Relation.Nullary using (¬_)
+
+------------------------------------------------------------------------
+-- Types
+
+infix 4 _∈_ _∉_
+
+_∈_ : A → ∀ {n} → Vec A n → Set _
+x ∈ xs = Any (x ≡_) xs
+
+_∉_ : A → ∀ {n} → Vec A n → Set _
+x ∉ xs = ¬ x ∈ xs
diff --git a/src/Data/Vec/Membership/Propositional/Properties.agda b/src/Data/Vec/Membership/Propositional/Properties.agda
new file mode 100644
index 0000000..07e3670
--- /dev/null
+++ b/src/Data/Vec/Membership/Propositional/Properties.agda
@@ -0,0 +1,103 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Properties of membership of vectors based on propositional equality.
+------------------------------------------------------------------------
+
+module Data.Vec.Membership.Propositional.Properties where
+
+open import Data.Fin using (Fin; zero; suc)
+open import Data.Product using (_,_)
+open import Data.Vec hiding (here; there)
+open import Data.Vec.Any using (here; there)
+open import Data.List using ([]; _∷_)
+open import Data.List.Any using (here; there)
+open import Data.Vec.Membership.Propositional
+import Data.List.Membership.Propositional as List
+open import Function using (_∘_; id)
+open import Relation.Binary.PropositionalEquality using (refl)
+
+------------------------------------------------------------------------
+-- lookup
+
+module _ {a} {A : Set a} where
+
+ ∈-lookup : ∀ {n} i (xs : Vec A n) → lookup i xs ∈ xs
+ ∈-lookup zero (x ∷ xs) = here refl
+ ∈-lookup (suc i) (x ∷ xs) = there (∈-lookup i xs)
+
+------------------------------------------------------------------------
+-- map
+
+module _ {a b} {A : Set a} {B : Set b} (f : A → B) where
+
+ ∈-map⁺ : ∀ {m v} {xs : Vec A m} → v ∈ xs → f v ∈ map f xs
+ ∈-map⁺ (here refl) = here refl
+ ∈-map⁺ (there x∈xs) = there (∈-map⁺ x∈xs)
+
+------------------------------------------------------------------------
+-- _++_
+
+module _ {a} {A : Set a} {v : A} where
+
+ ∈-++⁺ˡ : ∀ {m n} {xs : Vec A m} {ys : Vec A n} → v ∈ xs → v ∈ xs ++ ys
+ ∈-++⁺ˡ (here refl) = here refl
+ ∈-++⁺ˡ (there x∈xs) = there (∈-++⁺ˡ x∈xs)
+
+ ∈-++⁺ʳ : ∀ {m n} (xs : Vec A m) {ys : Vec A n} → v ∈ ys → v ∈ xs ++ ys
+ ∈-++⁺ʳ [] x∈ys = x∈ys
+ ∈-++⁺ʳ (x ∷ xs) x∈ys = there (∈-++⁺ʳ xs x∈ys)
+
+------------------------------------------------------------------------
+-- tabulate
+
+module _ {a} {A : Set a} where
+
+ ∈-tabulate⁺ : ∀ {n} (f : Fin n → A) i → f i ∈ tabulate f
+ ∈-tabulate⁺ f zero = here refl
+ ∈-tabulate⁺ f (suc i) = there (∈-tabulate⁺ (f ∘ suc) i)
+
+------------------------------------------------------------------------
+-- allFin
+
+∈-allFin⁺ : ∀ {n} (i : Fin n) → i ∈ allFin n
+∈-allFin⁺ = ∈-tabulate⁺ id
+
+------------------------------------------------------------------------
+-- allPairs
+
+module _ {a b} {A : Set a} {B : Set b} where
+
+ ∈-allPairs⁺ : ∀ {m n x y} {xs : Vec A m} {ys : Vec B n} →
+ x ∈ xs → y ∈ ys → (x , y) ∈ allPairs xs ys
+ ∈-allPairs⁺ {xs = x ∷ xs} (here refl) = ∈-++⁺ˡ ∘ ∈-map⁺ (x ,_)
+ ∈-allPairs⁺ {xs = x ∷ _} (there x∈xs) =
+ ∈-++⁺ʳ (map (x ,_) _) ∘ ∈-allPairs⁺ x∈xs
+
+------------------------------------------------------------------------
+-- toList
+
+module _ {a} {A : Set a} {v : A} where
+
+ ∈-toList⁺ : ∀ {n} {xs : Vec A n} → v ∈ xs → v List.∈ toList xs
+ ∈-toList⁺ (here refl) = here refl
+ ∈-toList⁺ (there x∈) = there (∈-toList⁺ x∈)
+
+ ∈-toList⁻ : ∀ {n} {xs : Vec A n} → v List.∈ toList xs → v ∈ xs
+ ∈-toList⁻ {xs = []} ()
+ ∈-toList⁻ {xs = x ∷ xs} (here refl) = here refl
+ ∈-toList⁻ {xs = x ∷ xs} (there v∈xs) = there (∈-toList⁻ v∈xs)
+
+------------------------------------------------------------------------
+-- fromList
+
+module _ {a} {A : Set a} {v : A} where
+
+ ∈-fromList⁺ : ∀ {xs} → v List.∈ xs → v ∈ fromList xs
+ ∈-fromList⁺ (here refl) = here refl
+ ∈-fromList⁺ (there x∈) = there (∈-fromList⁺ x∈)
+
+ ∈-fromList⁻ : ∀ {xs} → v ∈ fromList xs → v List.∈ xs
+ ∈-fromList⁻ {[]} ()
+ ∈-fromList⁻ {_ ∷ _} (here refl) = here refl
+ ∈-fromList⁻ {_ ∷ _} (there v∈xs) = there (∈-fromList⁻ v∈xs)
diff --git a/src/Data/Vec/Properties.agda b/src/Data/Vec/Properties.agda
index 8a3d370..27c1ff0 100644
--- a/src/Data/Vec/Properties.agda
+++ b/src/Data/Vec/Properties.agda
@@ -6,352 +6,436 @@
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)
-import Data.List.Any.Membership.Propositional as List
-open import Data.Nat
-open import Data.Nat.Properties using (+-assoc)
+open import Algebra.FunctionProperties
open import Data.Empty using (⊥-elim)
open import Data.Fin as Fin using (Fin; zero; suc; toℕ; fromℕ)
open import Data.Fin.Properties using (_+′_)
-open import Data.Product as Prod using (_×_; _,_; proj₁; proj₂; <_,_>)
+open import Data.List.Base as List using (List)
+open import Data.List.Any using (here; there)
+import Data.List.Membership.Propositional as List
+open import Data.Nat
+open import Data.Nat.Properties using (+-assoc; ≤-step)
+open import Data.Product as Prod
+ using (_×_; _,_; proj₁; proj₂; <_,_>; uncurry)
+open import Data.Vec
open import Function
-open import Function.Inverse using (_↔_)
-open import Relation.Binary
+open import Function.Inverse using (_↔_; inverse)
+open import Relation.Binary hiding (Decidable)
+open import Relation.Binary.PropositionalEquality as P
+ using (_≡_; _≢_; refl; _≗_)
+open import Relation.Binary.HeterogeneousEquality as H using (_≅_; refl)
+open import Relation.Unary using (Pred; Decidable)
+open import Relation.Nullary using (yes; no)
-module UsingVectorEquality {s₁ s₂} (S : Setoid s₁ s₂) where
+------------------------------------------------------------------------
+-- Properties of propositional equality over vectors
- private module SS = Setoid S
- open SS using () renaming (Carrier to A)
- import Data.Vec.Equality as VecEq
- open VecEq.Equality S
+module _ {a} {A : Set a} {n} {x y : A} {xs ys : Vec A n} where
- replicate-lemma : ∀ {m} n x (xs : Vec A m) →
- replicate {n = n} x ++ (x ∷ xs) ≈
- replicate {n = 1 + n} x ++ xs
- replicate-lemma zero x xs = refl (x ∷ xs)
- replicate-lemma (suc n) x xs = SS.refl ∷-cong replicate-lemma n x xs
+ ∷-injectiveˡ : x ∷ xs ≡ y ∷ ys → x ≡ y
+ ∷-injectiveˡ refl = refl
- xs++[]=xs : ∀ {n} (xs : Vec A n) → xs ++ [] ≈ xs
- xs++[]=xs [] = []-cong
- xs++[]=xs (x ∷ xs) = SS.refl ∷-cong xs++[]=xs xs
+ ∷-injectiveʳ : x ∷ xs ≡ y ∷ ys → xs ≡ ys
+ ∷-injectiveʳ refl = refl
- map-++-commute : ∀ {b m n} {B : Set b}
- (f : B → A) (xs : Vec B m) {ys : Vec B n} →
- map f (xs ++ ys) ≈ map f xs ++ map f ys
- map-++-commute f [] = refl _
- map-++-commute f (x ∷ xs) = SS.refl ∷-cong map-++-commute f xs
+ ∷-injective : (x ∷ xs) ≡ (y ∷ ys) → x ≡ y × xs ≡ ys
+ ∷-injective refl = refl , refl
-open import Relation.Binary.PropositionalEquality as P
- using (_≡_; _≢_; refl; _≗_)
-open import Relation.Binary.HeterogeneousEquality using (_≅_; refl)
+------------------------------------------------------------------------
+-- _[_]=_
-∷-injective : ∀ {a n} {A : Set a} {x y : A} {xs ys : Vec A n} →
- (x ∷ xs) ≡ (y ∷ ys) → x ≡ y × xs ≡ ys
-∷-injective refl = refl , refl
+module _ {a} {A : Set a} where
--- lookup is a functor morphism from Vec to Identity.
+ []=-injective : ∀ {n} {xs : Vec A n} {i x y} →
+ xs [ i ]= x → xs [ i ]= y → x ≡ y
+ []=-injective here here = refl
+ []=-injective (there xsᵢ≡x) (there xsᵢ≡y) = []=-injective xsᵢ≡x xsᵢ≡y
-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
+ []=-irrelevance : ∀ {n} {xs : Vec A n} {i x} →
+ (p q : xs [ i ]= x) → p ≡ q
+ []=-irrelevance here here = refl
+ []=-irrelevance (there xs[i]=x) (there xs[i]=x') =
+ P.cong there ([]=-irrelevance xs[i]=x xs[i]=x')
-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) →
- Morphism (applicative {a = a})
- (RawMonad.rawIApplicative IdentityMonad)
-lookup-morphism i = record
- { op = lookup i
- ; op-pure = lookup-replicate i
- ; op-⊛ = lookup-⊛ i
- }
- where
- 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
-
- lookup-⊛ : ∀ {a b n} {A : Set a} {B : Set b}
- i (fs : Vec (A → B) n) (xs : Vec A n) →
- lookup i (fs ⊛ xs) ≡ (lookup i fs $ lookup i xs)
- lookup-⊛ zero (f ∷ fs) (x ∷ xs) = refl
- lookup-⊛ (suc i) (f ∷ fs) (x ∷ xs) = lookup-⊛ i fs xs
+------------------------------------------------------------------------
+-- lookup
--- tabulate is an inverse of flip lookup.
+module _ {a} {A : Set a} where
-lookup∘tabulate : ∀ {a n} {A : Set a} (f : Fin n → A) (i : Fin n) →
- lookup i (tabulate f) ≡ f i
-lookup∘tabulate f zero = refl
-lookup∘tabulate f (suc i) = lookup∘tabulate (f ∘ suc) i
+ []=⇒lookup : ∀ {n} {x : A} {xs} {i : Fin n} →
+ xs [ i ]= x → lookup i xs ≡ x
+ []=⇒lookup here = refl
+ []=⇒lookup (there xs[i]=x) = []=⇒lookup xs[i]=x
-tabulate∘lookup : ∀ {a n} {A : Set a} (xs : Vec A n) →
- tabulate (flip lookup xs) ≡ xs
-tabulate∘lookup [] = refl
-tabulate∘lookup (x ∷ xs) = P.cong (_∷_ x) $ tabulate∘lookup xs
+ lookup⇒[]= : ∀ {n} (i : Fin n) {x : A} xs →
+ lookup i xs ≡ x → xs [ i ]= x
+ lookup⇒[]= zero (_ ∷ _) refl = here
+ lookup⇒[]= (suc i) (_ ∷ xs) p = there (lookup⇒[]= i xs p)
--- If you look up an index in allFin n, then you get the index.
+ []=↔lookup : ∀ {n i} {x} {xs : Vec A n} →
+ xs [ i ]= x ↔ lookup i xs ≡ x
+ []=↔lookup = inverse []=⇒lookup (lookup⇒[]= _ _)
+ (λ _ → []=-irrelevance _ _) (λ _ → P.≡-irrelevance _ _)
-lookup-allFin : ∀ {n} (i : Fin n) → lookup i (allFin n) ≡ i
-lookup-allFin = lookup∘tabulate id
+------------------------------------------------------------------------
+-- _[_]≔_ (update)
+
+module _ {a} {A : Set a} where
+
+ []≔-idempotent : ∀ {n} (xs : Vec A n) (i : Fin n) {x₁ x₂ : A} →
+ (xs [ i ]≔ x₁) [ i ]≔ x₂ ≡ xs [ i ]≔ x₂
+ []≔-idempotent [] ()
+ []≔-idempotent (x ∷ xs) zero = refl
+ []≔-idempotent (x ∷ xs) (suc i) = P.cong (x ∷_) ([]≔-idempotent xs i)
+
+ []≔-commutes : ∀ {n} (xs : Vec A n) (i j : Fin n) {x y : A} → i ≢ j →
+ (xs [ i ]≔ x) [ j ]≔ y ≡ (xs [ j ]≔ y) [ i ]≔ x
+ []≔-commutes [] () () _
+ []≔-commutes (x ∷ xs) zero zero 0≢0 = ⊥-elim $ 0≢0 refl
+ []≔-commutes (x ∷ xs) zero (suc i) _ = refl
+ []≔-commutes (x ∷ xs) (suc i) zero _ = refl
+ []≔-commutes (x ∷ xs) (suc i) (suc j) i≢j =
+ P.cong (x ∷_) $ []≔-commutes xs i j (i≢j ∘ P.cong suc)
+
+ []≔-updates : ∀ {n} (xs : Vec A n) (i : Fin n) {x : A} →
+ (xs [ i ]≔ x) [ i ]= x
+ []≔-updates [] ()
+ []≔-updates (x ∷ xs) zero = here
+ []≔-updates (x ∷ xs) (suc i) = there ([]≔-updates xs i)
+
+ []≔-minimal : ∀ {n} (xs : Vec A n) (i j : Fin n) {x y : A} → i ≢ j →
+ xs [ i ]= x → (xs [ j ]≔ y) [ i ]= x
+ []≔-minimal [] () () _ _
+ []≔-minimal (x ∷ xs) .zero zero 0≢0 here = ⊥-elim (0≢0 refl)
+ []≔-minimal (x ∷ xs) .zero (suc j) _ here = here
+ []≔-minimal (x ∷ xs) (suc i) zero _ (there loc) = there loc
+ []≔-minimal (x ∷ xs) (suc i) (suc j) i≢j (there loc) =
+ there ([]≔-minimal xs i j (i≢j ∘ P.cong suc) loc)
+
+ []≔-lookup : ∀ {n} (xs : Vec A n) (i : Fin n) →
+ xs [ i ]≔ lookup i xs ≡ xs
+ []≔-lookup [] ()
+ []≔-lookup (x ∷ xs) zero = refl
+ []≔-lookup (x ∷ xs) (suc i) = P.cong (_∷_ x) $ []≔-lookup xs i
+
+ lookup∘update : ∀ {n} (i : Fin n) (xs : Vec A n) x →
+ lookup i (xs [ i ]≔ x) ≡ x
+ lookup∘update zero (_ ∷ xs) x = refl
+ lookup∘update (suc i) (_ ∷ xs) x = lookup∘update i xs x
+
+ lookup∘update′ : ∀ {n} {i j : Fin n} → i ≢ j → ∀ (xs : Vec A n) y →
+ lookup i (xs [ j ]≔ y) ≡ lookup i xs
+ lookup∘update′ {i = zero} {zero} i≢j xs y = ⊥-elim (i≢j refl)
+ lookup∘update′ {i = zero} {suc j} i≢j (x ∷ xs) y = refl
+ lookup∘update′ {i = suc i} {zero} i≢j (x ∷ xs) y = refl
+ lookup∘update′ {i = suc i} {suc j} i≢j (x ∷ xs) y =
+ lookup∘update′ (i≢j ∘ P.cong suc) xs y
--- Various lemmas relating lookup and _++_.
-
-lookup-++-< : ∀ {a} {A : Set a} {m n}
- (xs : Vec A m) (ys : Vec A n) i (i<m : toℕ i < m) →
- lookup i (xs ++ ys) ≡ lookup (Fin.fromℕ≤ i<m) xs
-lookup-++-< [] ys i ()
-lookup-++-< (x ∷ xs) ys zero (s≤s z≤n) = refl
-lookup-++-< (x ∷ xs) ys (suc i) (s≤s (s≤s i<m)) =
- lookup-++-< xs ys i (s≤s i<m)
-
-lookup-++-≥ : ∀ {a} {A : Set a} {m n}
- (xs : Vec A m) (ys : Vec A n) i (i≥m : toℕ i ≥ m) →
- lookup i (xs ++ ys) ≡ lookup (Fin.reduce≥ i i≥m) ys
-lookup-++-≥ [] ys i i≥m = refl
-lookup-++-≥ (x ∷ xs) ys zero ()
-lookup-++-≥ (x ∷ xs) ys (suc i) (s≤s i≥m) = lookup-++-≥ xs ys i i≥m
-
-lookup-++-inject+ : ∀ {a} {A : Set a} {m n}
- (xs : Vec A m) (ys : Vec A n) i →
- lookup (Fin.inject+ n i) (xs ++ ys) ≡ lookup i xs
-lookup-++-inject+ [] ys ()
-lookup-++-inject+ (x ∷ xs) ys zero = refl
-lookup-++-inject+ (x ∷ xs) ys (suc i) = lookup-++-inject+ xs ys i
-
-lookup-++-+′ : ∀ {a} {A : Set a} {m n}
- (xs : Vec A m) (ys : Vec A n) i →
- lookup (fromℕ m +′ i) (xs ++ ys) ≡ lookup i ys
-lookup-++-+′ [] ys zero = refl
-lookup-++-+′ [] (y ∷ xs) (suc i) = lookup-++-+′ [] xs i
-lookup-++-+′ (x ∷ xs) ys i = lookup-++-+′ xs ys i
-
--- Properties relating lookup and _[_]≔_.
-
-lookup∘update : ∀ {a} {A : Set a} {n}
- (i : Fin n) (xs : Vec A n) x →
- lookup i (xs [ i ]≔ x) ≡ x
-lookup∘update zero (_ ∷ xs) x = refl
-lookup∘update (suc i) (_ ∷ xs) x = lookup∘update i xs x
-
-lookup∘update′ : ∀ {a} {A : Set a} {n} {i j : Fin n} →
- i ≢ j → ∀ (xs : Vec A n) y →
- lookup i (xs [ j ]≔ y) ≡ lookup i xs
-lookup∘update′ {i = zero} {zero} i≢j xs y = ⊥-elim (i≢j refl)
-lookup∘update′ {i = zero} {suc j} i≢j (x ∷ xs) y = refl
-lookup∘update′ {i = suc i} {zero} i≢j (x ∷ xs) y = refl
-lookup∘update′ {i = suc i} {suc j} i≢j (x ∷ xs) y =
- lookup∘update′ (i≢j ∘ P.cong suc) xs y
-
--- map is a congruence.
+------------------------------------------------------------------------
+-- map
+
+map-id : ∀ {a n} {A : Set a} → map {n = n} {A} id ≗ id
+map-id [] = refl
+map-id (x ∷ xs) = P.cong (x ∷_) (map-id xs)
map-cong : ∀ {a b n} {A : Set a} {B : Set b} {f g : A → B} →
- f ≗ g → _≗_ {A = Vec A n} (map f) (map g)
+ f ≗ g → map {n = n} 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 is functorial.
-
-map-id : ∀ {a n} {A : Set a} → _≗_ {A = Vec A n} (map id) id
-map-id [] = refl
-map-id (x ∷ xs) = P.cong (_∷_ x) (map-id xs)
-
map-∘ : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c}
(f : B → C) (g : A → B) →
- _≗_ {A = Vec A n} (map (f ∘ g)) (map f ∘ map g)
+ map {n = n} (f ∘ g) ≗ map f ∘ map g
map-∘ f g [] = refl
-map-∘ f g (x ∷ xs) = P.cong (_∷_ (f (g x))) (map-∘ f g xs)
-
--- 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.
+map-∘ f g (x ∷ xs) = P.cong (f (g x) ∷_) (map-∘ f g xs)
--- 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)
+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
--- pulling a replicate into a zipWith
+map-[]≔ : ∀ {n a b} {A : Set a} {B : Set b}
+ (f : A → B) (xs : Vec A n) (i : Fin n) {x : A} →
+ map f (xs [ i ]≔ x) ≡ map f xs [ i ]≔ f x
+map-[]≔ f [] ()
+map-[]≔ f (x ∷ xs) zero = refl
+map-[]≔ f (x ∷ xs) (suc i) = P.cong (_ ∷_) $ map-[]≔ f xs i
-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)
+------------------------------------------------------------------------
+-- _++_
+
+module _ {a} {A : Set a} {m} {ys ys' : Vec A m} where
+
+ ++-injectiveˡ : ∀ {n} (xs xs' : Vec A n) →
+ xs ++ ys ≡ xs' ++ ys' → xs ≡ xs'
+ ++-injectiveˡ [] [] _ = refl
+ ++-injectiveˡ (x ∷ xs) (x' ∷ xs') eq =
+ P.cong₂ _∷_ (∷-injectiveˡ eq) (++-injectiveˡ _ _ (∷-injectiveʳ eq))
+
+ ++-injectiveʳ : ∀ {n} (xs xs' : Vec A n) →
+ xs ++ ys ≡ xs' ++ ys' → ys ≡ ys'
+ ++-injectiveʳ [] [] eq = eq
+ ++-injectiveʳ (x ∷ xs) (x' ∷ xs') eq =
+ ++-injectiveʳ xs xs' (∷-injectiveʳ eq)
+
+ ++-injective : ∀ {n} (xs xs' : Vec A n) →
+ xs ++ ys ≡ xs' ++ ys' → xs ≡ xs' × ys ≡ ys'
+ ++-injective xs xs' eq =
+ (++-injectiveˡ xs xs' eq , ++-injectiveʳ xs xs' eq)
+
+module _ {a} {A : Set a} where
+
+ ++-assoc : ∀ {m n k} (xs : Vec A m) (ys : Vec A n) (zs : Vec A k) →
+ (xs ++ ys) ++ zs ≅ xs ++ (ys ++ zs)
+ ++-assoc [] ys zs = refl
+ ++-assoc {suc m} (x ∷ xs) ys zs =
+ H.icong (Vec A) (+-assoc m _ _) (x ∷_) (++-assoc xs ys zs)
+
+ lookup-++-< : ∀ {m n} (xs : Vec A m) (ys : Vec A n) →
+ ∀ i (i<m : toℕ i < m) →
+ lookup i (xs ++ ys) ≡ lookup (Fin.fromℕ≤ i<m) xs
+ lookup-++-< [] ys i ()
+ lookup-++-< (x ∷ xs) ys zero (s≤s z≤n) = refl
+ lookup-++-< (x ∷ xs) ys (suc i) (s≤s (s≤s i<m)) =
+ lookup-++-< xs ys i (s≤s i<m)
+
+ lookup-++-≥ : ∀ {m n} (xs : Vec A m) (ys : Vec A n) →
+ ∀ i (i≥m : toℕ i ≥ m) →
+ lookup i (xs ++ ys) ≡ lookup (Fin.reduce≥ i i≥m) ys
+ lookup-++-≥ [] ys i i≥m = refl
+ lookup-++-≥ (x ∷ xs) ys zero ()
+ lookup-++-≥ (x ∷ xs) ys (suc i) (s≤s i≥m) = lookup-++-≥ xs ys i i≥m
+
+ lookup-++-inject+ : ∀ {m n} (xs : Vec A m) (ys : Vec A n) i →
+ lookup (Fin.inject+ n i) (xs ++ ys) ≡ lookup i xs
+ lookup-++-inject+ [] ys ()
+ lookup-++-inject+ (x ∷ xs) ys zero = refl
+ lookup-++-inject+ (x ∷ xs) ys (suc i) = lookup-++-inject+ xs ys i
+
+ lookup-++-+′ : ∀ {m n} (xs : Vec A m) (ys : Vec A n) i →
+ lookup (fromℕ m +′ i) (xs ++ ys) ≡ lookup i ys
+ lookup-++-+′ [] ys zero = refl
+ lookup-++-+′ [] (y ∷ xs) (suc i) = lookup-++-+′ [] xs i
+ lookup-++-+′ (x ∷ xs) ys i = lookup-++-+′ xs ys i
+
+ []≔-++-inject+ : ∀ {m n x} (xs : Vec A m) (ys : Vec A n) i →
+ (xs ++ ys) [ Fin.inject+ n i ]≔ x ≡ (xs [ i ]≔ x) ++ ys
+ []≔-++-inject+ [] ys ()
+ []≔-++-inject+ (x ∷ xs) ys zero = refl
+ []≔-++-inject+ (x ∷ xs) ys (suc i) =
+ P.cong (x ∷_) $ []≔-++-inject+ xs ys i
-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)
+------------------------------------------------------------------------
+-- zipWith
+
+module _ {a} {A : Set a} {f : A → A → A} where
+
+ zipWith-assoc : Associative _≡_ f → ∀ {n} →
+ Associative _≡_ (zipWith {n = n} f)
+ zipWith-assoc assoc [] [] [] = refl
+ zipWith-assoc assoc (x ∷ xs) (y ∷ ys) (z ∷ zs) =
+ P.cong₂ _∷_ (assoc x y z) (zipWith-assoc assoc xs ys zs)
+
+ zipWith-idem : Idempotent _≡_ f → ∀ {n} →
+ Idempotent _≡_ (zipWith {n = n} f)
+ zipWith-idem idem [] = refl
+ zipWith-idem idem (x ∷ xs) =
+ P.cong₂ _∷_ (idem x) (zipWith-idem idem xs)
+
+ zipWith-identityˡ : ∀ {1#} → LeftIdentity _≡_ 1# f → ∀ {n} →
+ LeftIdentity _≡_ (replicate 1#) (zipWith {n = n} f)
+ zipWith-identityˡ idˡ [] = refl
+ zipWith-identityˡ idˡ (x ∷ xs) =
+ P.cong₂ _∷_ (idˡ x) (zipWith-identityˡ idˡ xs)
+
+ zipWith-identityʳ : ∀ {1#} → RightIdentity _≡_ 1# f → ∀ {n} →
+ RightIdentity _≡_ (replicate 1#) (zipWith {n = n} f)
+ zipWith-identityʳ idʳ [] = refl
+ zipWith-identityʳ idʳ (x ∷ xs) =
+ P.cong₂ _∷_ (idʳ x) (zipWith-identityʳ idʳ xs)
+
+ zipWith-zeroˡ : ∀ {0#} → LeftZero _≡_ 0# f → ∀ {n} →
+ LeftZero _≡_ (replicate 0#) (zipWith {n = n} f)
+ zipWith-zeroˡ zeˡ [] = refl
+ zipWith-zeroˡ zeˡ (x ∷ xs) =
+ P.cong₂ _∷_ (zeˡ x) (zipWith-zeroˡ zeˡ xs)
+
+ zipWith-zeroʳ : ∀ {0#} → RightZero _≡_ 0# f → ∀ {n} →
+ RightZero _≡_ (replicate 0#) (zipWith {n = n} f)
+ zipWith-zeroʳ zeʳ [] = refl
+ zipWith-zeroʳ zeʳ (x ∷ xs) =
+ P.cong₂ _∷_ (zeʳ x) (zipWith-zeroʳ zeʳ xs)
+
+ zipWith-inverseˡ : ∀ {⁻¹ 0#} → LeftInverse _≡_ 0# ⁻¹ f → ∀ {n} →
+ LeftInverse _≡_ (replicate {n = n} 0#) (map ⁻¹) (zipWith f)
+ zipWith-inverseˡ invˡ [] = refl
+ zipWith-inverseˡ invˡ (x ∷ xs) =
+ P.cong₂ _∷_ (invˡ x) (zipWith-inverseˡ invˡ xs)
+
+ zipWith-inverseʳ : ∀ {⁻¹ 0#} → RightInverse _≡_ 0# ⁻¹ f → ∀ {n} →
+ RightInverse _≡_ (replicate {n = n} 0#) (map ⁻¹) (zipWith f)
+ zipWith-inverseʳ invʳ [] = refl
+ zipWith-inverseʳ invʳ (x ∷ xs) =
+ P.cong₂ _∷_ (invʳ x) (zipWith-inverseʳ invʳ xs)
+
+ zipWith-distribˡ : ∀ {g} → _DistributesOverˡ_ _≡_ f g → ∀ {n} →
+ _DistributesOverˡ_ _≡_ (zipWith {n = n} f) (zipWith g)
+ zipWith-distribˡ distribˡ [] [] [] = refl
+ zipWith-distribˡ distribˡ (x ∷ xs) (y ∷ ys) (z ∷ zs) =
+ P.cong₂ _∷_ (distribˡ x y z) (zipWith-distribˡ distribˡ xs ys zs)
+
+ zipWith-distribʳ : ∀ {g} → _DistributesOverʳ_ _≡_ f g → ∀ {n} →
+ _DistributesOverʳ_ _≡_ (zipWith {n = n} f) (zipWith g)
+ zipWith-distribʳ distribʳ [] [] [] = refl
+ zipWith-distribʳ distribʳ (x ∷ xs) (y ∷ ys) (z ∷ zs) =
+ P.cong₂ _∷_ (distribʳ x y z) (zipWith-distribʳ distribʳ xs ys zs)
+
+ zipWith-absorbs : ∀ {g} → _Absorbs_ _≡_ f g → ∀ {n} →
+ _Absorbs_ _≡_ (zipWith {n = n} f) (zipWith g)
+ zipWith-absorbs abs [] [] = refl
+ zipWith-absorbs abs (x ∷ xs) (y ∷ ys) =
+ P.cong₂ _∷_ (abs x y) (zipWith-absorbs abs xs ys)
+
+module _ {a b} {A : Set a} {B : Set b} {f : A → A → B} where
+
+ zipWith-comm : (∀ x y → f x y ≡ f y x) → ∀ {n}
+ (xs ys : Vec A n) → zipWith f xs ys ≡ zipWith f ys xs
+ zipWith-comm comm [] [] = refl
+ zipWith-comm comm (x ∷ xs) (y ∷ ys) =
+ P.cong₂ _∷_ (comm x y) (zipWith-comm comm xs ys)
+
+module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
+
+ zipWith-map₁ : ∀ {n} (_⊕_ : 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₂ : ∀ {n} (_⊕_ : 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)
+
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
+
+ lookup-zipWith : ∀ (f : A → B → C) {n} (i : Fin n) xs ys →
+ lookup i (zipWith f xs ys) ≡ f (lookup i xs) (lookup i ys)
+ lookup-zipWith _ zero (x ∷ _) (y ∷ _) = refl
+ lookup-zipWith _ (suc i) (_ ∷ xs) (_ ∷ ys) = lookup-zipWith _ i xs ys
--- pulling a map into a zipWith
+------------------------------------------------------------------------
+-- zip
-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)
+module _ {a b} {A : Set a} {B : Set b} where
-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)
+ lookup-zip : ∀ {n} (i : Fin n) (xs : Vec A n) (ys : Vec B n) →
+ lookup i (zip xs ys) ≡ (lookup i xs , lookup i ys)
+ lookup-zip = lookup-zipWith _,_
--- Tabulation.
+ -- map lifts projections to vectors of products.
--- mapping over a tabulation is the tabulation of the composition
+ map-proj₁-zip : ∀ {n} (xs : Vec A n) (ys : Vec B n) →
+ map proj₁ (zip xs ys) ≡ xs
+ map-proj₁-zip [] [] = refl
+ map-proj₁-zip (x ∷ xs) (y ∷ ys) = P.cong (x ∷_) (map-proj₁-zip xs ys)
-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))
+ map-proj₂-zip : ∀ {n} (xs : Vec A n) (ys : Vec B n) →
+ map proj₂ (zip xs ys) ≡ ys
+ map-proj₂-zip [] [] = refl
+ map-proj₂-zip (x ∷ xs) (y ∷ ys) = P.cong (y ∷_) (map-proj₂-zip xs ys)
--- tabulate can be expressed using map and allFin.
+-- map lifts pairing to vectors of products.
-tabulate-allFin : ∀ {n a} {A : Set a} (f : Fin n → A) →
- tabulate f ≡ map f (allFin n)
-tabulate-allFin f = tabulate-∘ f id
+map-<,>-zip : ∀ {a b c n} {A : Set a} {B : Set b} {C : Set c}
+ (f : A → B) (g : A → C) (xs : Vec A n) →
+ map < f , g > xs ≡ zip (map f xs) (map g xs)
+map-<,>-zip f g [] = P.refl
+map-<,>-zip f g (x ∷ xs) = P.cong (_ ∷_) (map-<,>-zip f g xs)
--- The positive case of allFin can be expressed recursively using map.
+map-zip : ∀ {a b c d n} {A : Set a} {B : Set b} {C : Set c} {D : Set d}
+ (f : A → B) (g : C → D) (xs : Vec A n) (ys : Vec C n) →
+ map (Prod.map f g) (zip xs ys) ≡ zip (map f xs) (map g ys)
+map-zip f g [] [] = refl
+map-zip f g (x ∷ xs) (y ∷ ys) = P.cong (_ ∷_) (map-zip f g xs ys)
-allFin-map : ∀ n → allFin (suc n) ≡ zero ∷ map suc (allFin n)
-allFin-map n = P.cong (_∷_ zero) $ tabulate-∘ suc id
+------------------------------------------------------------------------
+-- unzip
--- If you look up every possible index, in increasing order, then you
--- get back the vector you started with.
+module _ {a b} {A : Set a} {B : Set b} where
-map-lookup-allFin : ∀ {a} {A : Set a} {n} (xs : Vec A n) →
- map (λ x → lookup x xs) (allFin n) ≡ xs
-map-lookup-allFin {n = n} xs = begin
- map (λ x → lookup x xs) (allFin n) ≡⟨ P.sym $ tabulate-∘ (λ x → lookup x xs) id ⟩
- tabulate (λ x → lookup x xs) ≡⟨ tabulate∘lookup xs ⟩
- xs ∎
- where open P.≡-Reasoning
+ lookup-unzip : ∀ {n} (i : Fin n) (xys : Vec (A × B) n) →
+ let xs , ys = unzip xys
+ in (lookup i xs , lookup i ys) ≡ lookup i xys
+ lookup-unzip () []
+ lookup-unzip zero ((x , y) ∷ xys) = refl
+ lookup-unzip (suc i) ((x , y) ∷ xys) = lookup-unzip i xys
--- tabulate f contains f i.
+ map-unzip : ∀ {c d n} {C : Set c} {D : Set d}
+ (f : A → B) (g : C → D) (xys : Vec (A × C) n) →
+ let xs , ys = unzip xys
+ in (map f xs , map g ys) ≡ unzip (map (Prod.map f g) xys)
+ map-unzip f g [] = refl
+ map-unzip f g ((x , y) ∷ xys) =
+ P.cong (Prod.map (f x ∷_) (g y ∷_)) (map-unzip f g xys)
-∈-tabulate : ∀ {n a} {A : Set a} (f : Fin n → A) i → f i ∈ tabulate f
-∈-tabulate f zero = here
-∈-tabulate f (suc i) = there (∈-tabulate (f ∘ suc) i)
+ -- Products of vectors are isomorphic to vectors of products.
--- allFin n contains all elements in Fin n.
+ unzip∘zip : ∀ {n} (xs : Vec A n) (ys : Vec B n) →
+ unzip (zip xs ys) ≡ (xs , ys)
+ unzip∘zip [] [] = refl
+ unzip∘zip (x ∷ xs) (y ∷ ys) =
+ P.cong (Prod.map (x ∷_) (y ∷_)) (unzip∘zip xs ys)
-∈-allFin : ∀ {n} (i : Fin n) → i ∈ allFin n
-∈-allFin = ∈-tabulate id
+ zip∘unzip : ∀ {n} (xys : Vec (A × B) n) →
+ uncurry zip (unzip xys) ≡ xys
+ zip∘unzip [] = refl
+ zip∘unzip ((x , y) ∷ xys) = P.cong ((x , y) ∷_) (zip∘unzip xys)
-∈-map : ∀ {a b m} {A : Set a} {B : Set b} {x : A} {xs : Vec A m}
- (f : A → B) → x ∈ xs → f x ∈ map f xs
-∈-map f here = here
-∈-map f (there x∈xs) = there (∈-map f x∈xs)
+ ×v↔v× : ∀ {n} → (Vec A n × Vec B n) ↔ Vec (A × B) n
+ ×v↔v× = inverse (uncurry zip) unzip (uncurry unzip∘zip) zip∘unzip
--- _∈_ is preserved by _++_
-∈-++ₗ : ∀ {a m n} {A : Set a} {x : A} {xs : Vec A m} {ys : Vec A n}
- → x ∈ xs → x ∈ xs ++ ys
-∈-++ₗ here = here
-∈-++ₗ (there x∈xs) = there (∈-++ₗ x∈xs)
+------------------------------------------------------------------------
+-- _⊛_
-∈-++ᵣ : ∀ {a m n} {A : Set a} {x : A} (xs : Vec A m) {ys : Vec A n}
- → x ∈ ys → x ∈ xs ++ ys
-∈-++ᵣ [] x∈ys = x∈ys
-∈-++ᵣ (x ∷ xs) x∈ys = there (∈-++ᵣ xs x∈ys)
+module _ {a b} {A : Set a} {B : Set b} where
--- allPairs contains every pair of elements
+ lookup-⊛ : ∀ {n} i (fs : Vec (A → B) n) (xs : Vec A n) →
+ lookup i (fs ⊛ xs) ≡ (lookup i fs $ lookup i xs)
+ lookup-⊛ zero (f ∷ fs) (x ∷ xs) = refl
+ lookup-⊛ (suc i) (f ∷ fs) (x ∷ xs) = lookup-⊛ i fs xs
-∈-allPairs : ∀ {a b} {A : Set a} {B : Set b} {m n : ℕ}
- {xs : Vec A m} {ys : Vec B n}
- {x y} → x ∈ xs → y ∈ ys → (x , y) ∈ allPairs xs ys
-∈-allPairs {xs = x ∷ xs} {ys} here y∈ys =
- ∈-++ₗ (∈-map (x ,_) y∈ys)
-∈-allPairs {xs = x ∷ xs} {ys} (there x∈xs) y∈ys =
- ∈-++ᵣ (map (x ,_) ys) (∈-allPairs x∈xs y∈ys)
+ map-is-⊛ : ∀ {n} (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)
--- sum commutes with _++_.
+ ⊛-is-zipWith : ∀ {n} (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)
-sum-++-commute : ∀ {m n} (xs : Vec ℕ m) {ys : Vec ℕ n} →
- sum (xs ++ ys) ≡ sum xs + sum ys
-sum-++-commute [] = refl
-sum-++-commute (x ∷ xs) {ys} = begin
- x + sum (xs ++ ys)
- ≡⟨ P.cong (λ p → x + p) (sum-++-commute xs) ⟩
- x + (sum xs + sum ys)
- ≡⟨ P.sym (+-assoc x (sum xs) (sum ys)) ⟩
- sum (x ∷ xs) + sum ys
- ∎
- where
- open P.≡-Reasoning
+ zipWith-is-⊛ : ∀ {c} {C : Set c} {n} (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)
--- foldr is a congruence.
+------------------------------------------------------------------------
+-- foldr
foldr-cong : ∀ {a b} {A : Set a}
- {B₁ : ℕ → Set b}
- {f₁ : ∀ {n} → A → B₁ n → B₁ (suc n)} {e₁}
- {B₂ : ℕ → Set b}
- {f₂ : ∀ {n} → A → B₂ n → B₂ (suc n)} {e₂} →
- (∀ {n x} {y₁ : B₁ n} {y₂ : B₂ n} →
- y₁ ≅ y₂ → f₁ x y₁ ≅ f₂ x y₂) →
- e₁ ≅ e₂ →
- ∀ {n} (xs : Vec A n) →
- foldr B₁ f₁ e₁ xs ≅ foldr B₂ f₂ e₂ xs
-foldr-cong _ e₁=e₂ [] = e₁=e₂
-foldr-cong {B₁ = B₁} f₁=f₂ e₁=e₂ (x ∷ xs) =
- f₁=f₂ (foldr-cong {B₁ = B₁} f₁=f₂ e₁=e₂ xs)
-
--- foldl is a congruence.
-
-foldl-cong : ∀ {a b} {A : Set a}
- {B₁ : ℕ → Set b}
- {f₁ : ∀ {n} → B₁ n → A → B₁ (suc n)} {e₁}
- {B₂ : ℕ → Set b}
- {f₂ : ∀ {n} → B₂ n → A → B₂ (suc n)} {e₂} →
- (∀ {n x} {y₁ : B₁ n} {y₂ : B₂ n} →
- y₁ ≅ y₂ → f₁ y₁ x ≅ f₂ y₂ x) →
- e₁ ≅ e₂ →
- ∀ {n} (xs : Vec A n) →
- foldl B₁ f₁ e₁ xs ≅ foldl B₂ f₂ e₂ xs
-foldl-cong _ e₁=e₂ [] = e₁=e₂
-foldl-cong {B₁ = B₁} f₁=f₂ e₁=e₂ (x ∷ xs) =
- foldl-cong {B₁ = B₁ ∘ suc} f₁=f₂ (f₁=f₂ e₁=e₂) xs
+ {B : ℕ → Set b} {f : ∀ {n} → A → B n → B (suc n)} {d}
+ {C : ℕ → Set b} {g : ∀ {n} → A → C n → C (suc n)} {e} →
+ (∀ {n x} {y : B n} {z : C n} → y ≅ z → f x y ≅ g x z) →
+ d ≅ e → ∀ {n} (xs : Vec A n) →
+ foldr B f d xs ≅ foldr C g e xs
+foldr-cong _ d≅e [] = d≅e
+foldr-cong f≅g d≅e (x ∷ xs) = f≅g (foldr-cong f≅g d≅e xs)
-- The (uniqueness part of the) universality property for foldr.
@@ -359,9 +443,9 @@ foldr-universal : ∀ {a b} {A : Set a} (B : ℕ → Set b)
(f : ∀ {n} → A → B n → B (suc n)) {e}
(h : ∀ {n} → Vec A n → B n) →
h [] ≡ e →
- (∀ {n} x → h ∘ _∷_ x ≗ f {n} x ∘ h) →
+ (∀ {n} x → h ∘ (x ∷_) ≗ f {n} x ∘ h) →
∀ {n} → h ≗ foldr B {n} f e
-foldr-universal B f h base step [] = base
+foldr-universal B f {_} h base step [] = base
foldr-universal B f {e} h base step (x ∷ xs) = begin
h (x ∷ xs)
≡⟨ step x xs ⟩
@@ -371,199 +455,194 @@ foldr-universal B f {e} h base step (x ∷ xs) = begin
where open P.≡-Reasoning
--- A fusion law for foldr.
-
foldr-fusion : ∀ {a b c} {A : Set a}
- {B : ℕ → Set b} {f : ∀ {n} → A → B n → B (suc n)} e
- {C : ℕ → Set c} {g : ∀ {n} → A → C n → C (suc n)}
+ {B : ℕ → Set b} {f : ∀ {n} → A → B n → B (suc n)} e
+ {C : ℕ → Set c} {g : ∀ {n} → A → C n → C (suc n)}
(h : ∀ {n} → B n → C n) →
(∀ {n} x → h ∘ f {n} x ≗ g x ∘ h) →
∀ {n} → h ∘ foldr B {n} f e ≗ foldr C g (h e)
foldr-fusion {B = B} {f} e {C} h fuse =
foldr-universal C _ _ refl (λ x xs → fuse x (foldr B f e xs))
--- The identity function is a fold.
-
idIsFold : ∀ {a n} {A : Set a} → id ≗ foldr (Vec A) {n} _∷_ []
idIsFold = foldr-universal _ _ id refl (λ _ _ → refl)
--- The _∈_ predicate is equivalent (in the following sense) to the
--- corresponding predicate for lists.
-
-∈⇒List-∈ : ∀ {a} {A : Set a} {n x} {xs : Vec A n} →
- x ∈ xs → x List.∈ toList xs
-∈⇒List-∈ here = here P.refl
-∈⇒List-∈ (there x∈) = there (∈⇒List-∈ x∈)
-
-List-∈⇒∈ : ∀ {a} {A : Set a} {x : A} {xs} →
- x List.∈ xs → x ∈ fromList xs
-List-∈⇒∈ (here P.refl) = here
-List-∈⇒∈ (there x∈) = there (List-∈⇒∈ x∈)
-
--- Proof irrelevance for _[_]=_.
-
-proof-irrelevance-[]= : ∀ {a} {A : Set a} {n} {xs : Vec A n} {i x} →
- (p q : xs [ i ]= x) → p ≡ q
-proof-irrelevance-[]= here here = refl
-proof-irrelevance-[]= (there xs[i]=x) (there xs[i]=x') =
- P.cong there (proof-irrelevance-[]= xs[i]=x xs[i]=x')
-
--- _[_]=_ can be expressed using lookup and _≡_.
-
-[]=↔lookup : ∀ {a n i} {A : Set a} {x} {xs : Vec A n} →
- xs [ i ]= x ↔ lookup i xs ≡ x
-[]=↔lookup {i = i} {x = x} {xs} = record
- { to = P.→-to-⟶ to
- ; from = P.→-to-⟶ (from i xs)
- ; inverse-of = record
- { left-inverse-of = λ _ → proof-irrelevance-[]= _ _
- ; right-inverse-of = λ _ → P.proof-irrelevance _ _
- }
- }
- where
- to : ∀ {n xs} {i : Fin n} → xs [ i ]= x → lookup i xs ≡ x
- to here = refl
- to (there xs[i]=x) = to xs[i]=x
-
- from : ∀ {n} (i : Fin n) xs → lookup i xs ≡ x → xs [ i ]= x
- from zero (.x ∷ _) refl = here
- from (suc i) (_ ∷ xs) p = there (from i xs p)
+------------------------------------------------------------------------
+-- foldl
+
+foldl-cong : ∀ {a b} {A : Set a}
+ {B : ℕ → Set b} {f : ∀ {n} → B n → A → B (suc n)} {d}
+ {C : ℕ → Set b} {g : ∀ {n} → C n → A → C (suc n)} {e} →
+ (∀ {n x} {y : B n} {z : C n} → y ≅ z → f y x ≅ g z x) →
+ d ≅ e → ∀ {n} (xs : Vec A n) →
+ foldl B f d xs ≅ foldl C g e xs
+foldl-cong _ d≅e [] = d≅e
+foldl-cong f≅g d≅e (x ∷ xs) = foldl-cong f≅g (f≅g d≅e) xs
------------------------------------------------------------------------
--- Some properties related to _[_]≔_
-
-[]≔-idempotent :
- ∀ {n a} {A : Set a} (xs : Vec A n) (i : Fin n) {x₁ x₂ : A} →
- (xs [ i ]≔ x₁) [ i ]≔ x₂ ≡ xs [ i ]≔ x₂
-[]≔-idempotent [] ()
-[]≔-idempotent (x ∷ xs) zero = refl
-[]≔-idempotent (x ∷ xs) (suc i) = P.cong (_∷_ x) $ []≔-idempotent xs i
-
-[]≔-commutes :
- ∀ {n a} {A : Set a} (xs : Vec A n) (i j : Fin n) {x y : A} →
- i ≢ j → (xs [ i ]≔ x) [ j ]≔ y ≡ (xs [ j ]≔ y) [ i ]≔ x
-[]≔-commutes [] () () _
-[]≔-commutes (x ∷ xs) zero zero 0≢0 = ⊥-elim $ 0≢0 refl
-[]≔-commutes (x ∷ xs) zero (suc i) _ = refl
-[]≔-commutes (x ∷ xs) (suc i) zero _ = refl
-[]≔-commutes (x ∷ xs) (suc i) (suc j) i≢j =
- P.cong (_∷_ x) $ []≔-commutes xs i j (i≢j ∘ P.cong suc)
-
-[]≔-updates : ∀ {n a} {A : Set a} (xs : Vec A n) (i : Fin n) {x : A} →
- (xs [ i ]≔ x) [ i ]= x
-[]≔-updates [] ()
-[]≔-updates (x ∷ xs) zero = here
-[]≔-updates (x ∷ xs) (suc i) = there ([]≔-updates xs i)
-
-[]≔-minimal :
- ∀ {n a} {A : Set a} (xs : Vec A n) (i j : Fin n) {x y : A} →
- i ≢ j → xs [ i ]= x → (xs [ j ]≔ y) [ i ]= x
-[]≔-minimal [] () () _ _
-[]≔-minimal (x ∷ xs) .zero zero 0≢0 here = ⊥-elim $ 0≢0 refl
-[]≔-minimal (x ∷ xs) .zero (suc j) _ here = here
-[]≔-minimal (x ∷ xs) (suc i) zero _ (there loc) = there loc
-[]≔-minimal (x ∷ xs) (suc i) (suc j) i≢j (there loc) =
- there ([]≔-minimal xs i j (i≢j ∘ P.cong suc) loc)
+-- sum
-map-[]≔ : ∀ {n a b} {A : Set a} {B : Set b}
- (f : A → B) (xs : Vec A n) (i : Fin n) {x : A} →
- map f (xs [ i ]≔ x) ≡ map f xs [ i ]≔ f x
-map-[]≔ f [] ()
-map-[]≔ f (x ∷ xs) zero = refl
-map-[]≔ f (x ∷ xs) (suc i) = P.cong (_∷_ _) $ map-[]≔ f xs i
-
-[]≔-lookup : ∀ {a} {A : Set a} {n} (xs : Vec A n) (i : Fin n) →
- xs [ i ]≔ lookup i xs ≡ xs
-[]≔-lookup [] ()
-[]≔-lookup (x ∷ xs) zero = refl
-[]≔-lookup (x ∷ xs) (suc i) = P.cong (_∷_ x) $ []≔-lookup xs i
-
-[]≔-++-inject+ : ∀ {a} {A : Set a} {m n x}
- (xs : Vec A m) (ys : Vec A n) i →
- (xs ++ ys) [ Fin.inject+ n i ]≔ x ≡ xs [ i ]≔ x ++ ys
-[]≔-++-inject+ [] ys ()
-[]≔-++-inject+ (x ∷ xs) ys zero = refl
-[]≔-++-inject+ (x ∷ xs) ys (suc i) =
- P.cong (_∷_ x) $ []≔-++-inject+ xs ys i
+sum-++-commute : ∀ {m n} (xs : Vec ℕ m) {ys : Vec ℕ n} →
+ sum (xs ++ ys) ≡ sum xs + sum ys
+sum-++-commute [] {_} = refl
+sum-++-commute (x ∷ xs) {ys} = begin
+ x + sum (xs ++ ys) ≡⟨ P.cong (x +_) (sum-++-commute xs) ⟩
+ x + (sum xs + sum ys) ≡⟨ P.sym (+-assoc x (sum xs) (sum ys)) ⟩
+ sum (x ∷ xs) + sum ys ∎
+ where open P.≡-Reasoning
------------------------------------------------------------------------
--- Some properties related to zipping and unzipping.
-
--- Products of vectors are isomorphic to vectors of products.
-
-unzip∘zip : ∀ {a n} {A B : Set a} (xs : Vec A n) (ys : Vec B n) →
- unzip (zip xs ys) ≡ (xs , ys)
-unzip∘zip [] [] = refl
-unzip∘zip (x ∷ xs) (y ∷ ys) =
- P.cong (Prod.map (_∷_ x) (_∷_ y)) (unzip∘zip xs ys)
-
-zip∘unzip : ∀ {a n} {A B : Set a} (xys : Vec (A × B) n) →
- (Prod.uncurry zip) (unzip xys) ≡ xys
-zip∘unzip [] = refl
-zip∘unzip ((x , y) ∷ xys) = P.cong (_∷_ (x , y)) (zip∘unzip xys)
-
-×v↔v× : ∀ {a n} {A B : Set a} → (Vec A n × Vec B n) ↔ Vec (A × B) n
-×v↔v× = record
- { to = P.→-to-⟶ (Prod.uncurry zip)
- ; from = P.→-to-⟶ unzip
- ; inverse-of = record
- { left-inverse-of = Prod.uncurry unzip∘zip
- ; right-inverse-of = zip∘unzip
- }
- }
-
--- map lifts projections to vectors of products.
-
-map-proj₁-zip : ∀ {a n} {A B : Set a} (xs : Vec A n) (ys : Vec B n) →
- map proj₁ (zip xs ys) ≡ xs
-map-proj₁-zip [] [] = refl
-map-proj₁-zip (x ∷ xs) (y ∷ ys) = P.cong (_∷_ x) (map-proj₁-zip xs ys)
-
-map-proj₂-zip : ∀ {a n} {A B : Set a} (xs : Vec A n) (ys : Vec B n) →
- map proj₂ (zip xs ys) ≡ ys
-map-proj₂-zip [] [] = refl
-map-proj₂-zip (x ∷ xs) (y ∷ ys) = P.cong (_∷_ y) (map-proj₂-zip xs ys)
+-- replicate
--- map lifts pairing to vectors of products.
+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
-map-<,>-zip : ∀ {a n} {A B₁ B₂ : Set a}
- (f : A → B₁) (g : A → B₂) (xs : Vec A n) →
- map < f , g > xs ≡ zip (map f xs) (map g xs)
-map-<,>-zip f g [] = P.refl
-map-<,>-zip f g (x ∷ xs) = P.cong (_∷_ (f x , g x)) (map-<,>-zip f g xs)
+map-replicate : ∀ {a b} {A : Set a} {B : Set b} (f : A → B) (x : A) →
+ ∀ n → map f (replicate 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)
-map-zip : ∀ {a n} {A₁ A₂ B₁ B₂ : Set a}
- (f : A₁ → A₂) (g : B₁ → B₂) (xs : Vec A₁ n) (ys : Vec B₁ n) →
- map (Prod.map f g) (zip xs ys) ≡ zip (map f xs) (map g ys)
-map-zip f g [] [] = refl
-map-zip f g (x ∷ xs) (y ∷ ys) = P.cong (_∷_ (f x , g y)) (map-zip f g xs ys)
-
-map-unzip : ∀ {a n} {A₁ A₂ B₁ B₂ : Set a}
- (f : A₁ → A₂) (g : B₁ → B₂) (xys : Vec (A₁ × B₁) n) →
- let xs , ys = unzip xys
- in (map f xs , map g ys) ≡ unzip (map (Prod.map f g) xys)
-map-unzip f g [] = refl
-map-unzip f g ((x , y) ∷ xys) =
- P.cong (Prod.map (_∷_ (f x)) (_∷_ (g y))) (map-unzip f g xys)
-
--- lookup is homomorphic with respect to the product structure.
-
-lookup-unzip : ∀ {a n} {A B : Set a} (i : Fin n) (xys : Vec (A × B) n) →
- let xs , ys = unzip xys
- in (lookup i xs , lookup i ys) ≡ lookup i xys
-lookup-unzip () []
-lookup-unzip zero ((x , y) ∷ xys) = refl
-lookup-unzip (suc i) ((x , y) ∷ xys) = lookup-unzip i xys
-
-lookup-zip : ∀ {a n} {A B : Set a}
- (i : Fin n) (xs : Vec A n) (ys : Vec B n) →
- lookup i (zip xs ys) ≡ (lookup i xs , lookup i ys)
-lookup-zip i xs ys = begin
- lookup i (zip xs ys)
- ≡⟨ P.sym (lookup-unzip i (zip xs ys)) ⟩
- lookup i (proj₁ (unzip (zip xs ys))) , lookup i (proj₂ (unzip (zip xs ys)))
- ≡⟨ P.cong₂ _,_ (P.cong (lookup i ∘ proj₁) (unzip∘zip xs ys))
- (P.cong (lookup i ∘ proj₂) (unzip∘zip xs ys)) ⟩
- lookup i xs , lookup i ys
- ∎
+module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where
+
+ zipWith-replicate₁ : ∀ {n} (_⊕_ : 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₂ : ∀ {n} (_⊕_ : 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)
+
+------------------------------------------------------------------------
+-- tabulate
+
+lookup∘tabulate : ∀ {a n} {A : Set a} (f : Fin n → A) (i : Fin n) →
+ lookup i (tabulate f) ≡ f i
+lookup∘tabulate f zero = refl
+lookup∘tabulate f (suc i) = lookup∘tabulate (f ∘ suc) i
+
+tabulate∘lookup : ∀ {a n} {A : Set a} (xs : Vec A n) →
+ tabulate (flip lookup xs) ≡ xs
+tabulate∘lookup [] = refl
+tabulate∘lookup (x ∷ xs) = P.cong (x ∷_) (tabulate∘lookup xs)
+
+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-cong : ∀ {n a} {A : Set a} {f g : Fin n → A} → f ≗ g → tabulate f ≡ tabulate g
+tabulate-cong {zero} p = refl
+tabulate-cong {suc n} p = P.cong₂ _∷_ (p zero) (tabulate-cong (p ∘ suc))
+
+------------------------------------------------------------------------
+-- allFin
+
+lookup-allFin : ∀ {n} (i : Fin n) → lookup i (allFin n) ≡ i
+lookup-allFin = lookup∘tabulate id
+
+allFin-map : ∀ n → allFin (suc n) ≡ zero ∷ map suc (allFin n)
+allFin-map n = P.cong (zero ∷_) $ tabulate-∘ suc id
+
+tabulate-allFin : ∀ {n a} {A : Set a} (f : Fin n → A) →
+ tabulate f ≡ map f (allFin n)
+tabulate-allFin f = tabulate-∘ f id
+
+-- If you look up every possible index, in increasing order, then you
+-- get back the vector you started with.
+
+map-lookup-allFin : ∀ {a} {A : Set a} {n} (xs : Vec A n) →
+ map (λ x → lookup x xs) (allFin n) ≡ xs
+map-lookup-allFin {n = n} xs = begin
+ map (λ x → lookup x xs) (allFin n) ≡⟨ P.sym $ tabulate-∘ (λ x → lookup x xs) id ⟩
+ tabulate (λ x → lookup x xs) ≡⟨ tabulate∘lookup xs ⟩
+ xs ∎
where open P.≡-Reasoning
+
+------------------------------------------------------------------------
+-- count
+
+module _ {a p} {A : Set a} {P : Pred A p} (P? : Decidable P) where
+
+ count≤n : ∀ {n} (xs : Vec A n) → count P? xs ≤ n
+ count≤n [] = z≤n
+ count≤n (x ∷ xs) with P? x
+ ... | yes _ = s≤s (count≤n xs)
+ ... | no _ = ≤-step (count≤n xs)
+
+------------------------------------------------------------------------
+-- insert
+
+module _ {a} {A : Set a} where
+
+ insert-lookup : ∀ {n} (i : Fin (suc n)) (x : A)
+ (xs : Vec A n) → lookup i (insert i x xs) ≡ x
+ insert-lookup zero x xs = refl
+ insert-lookup (suc ()) x []
+ insert-lookup (suc i) x (y ∷ xs) = insert-lookup i x xs
+
+ insert-punchIn : ∀ {n} (i : Fin (suc n)) (x : A) (xs : Vec A n)
+ (j : Fin n) →
+ lookup (Fin.punchIn i j) (insert i x xs) ≡ lookup j xs
+ insert-punchIn zero x xs j = refl
+ insert-punchIn (suc ()) x [] j
+ insert-punchIn (suc i) x (y ∷ xs) zero = refl
+ insert-punchIn (suc i) x (y ∷ xs) (suc j) = insert-punchIn i x xs j
+
+ remove-punchOut : ∀ {n} (xs : Vec A (suc n))
+ {i : Fin (suc n)} {j : Fin (suc n)} (i≢j : i ≢ j) →
+ lookup (Fin.punchOut i≢j) (remove i xs) ≡ lookup j xs
+ remove-punchOut (x ∷ xs) {zero} {zero} i≢j = ⊥-elim (i≢j refl)
+ remove-punchOut (x ∷ xs) {zero} {suc j} i≢j = refl
+ remove-punchOut (x ∷ []) {suc ()} {j} i≢j
+ remove-punchOut (x ∷ y ∷ xs) {suc i} {zero} i≢j = refl
+ remove-punchOut (x ∷ y ∷ xs) {suc i} {suc j} i≢j =
+ remove-punchOut (y ∷ xs) (i≢j ∘ P.cong suc)
+
+------------------------------------------------------------------------
+-- remove
+
+ remove-insert : ∀ {n} (i : Fin (suc n)) (x : A) (xs : Vec A n) →
+ remove i (insert i x xs) ≡ xs
+ remove-insert zero x xs = refl
+ remove-insert (suc ()) x []
+ remove-insert (suc zero) x (y ∷ xs) = refl
+ remove-insert (suc (suc ())) x (y ∷ [])
+ remove-insert (suc (suc i)) x (y ∷ z ∷ xs) =
+ P.cong (y ∷_) (remove-insert (suc i) x (z ∷ xs))
+
+ insert-remove : ∀ {n} (i : Fin (suc n)) (xs : Vec A (suc n)) →
+ insert i (lookup i xs) (remove i xs) ≡ xs
+ insert-remove zero (x ∷ xs) = refl
+ insert-remove (suc ()) (x ∷ [])
+ insert-remove (suc i) (x ∷ y ∷ xs) =
+ P.cong (x ∷_) (insert-remove i (y ∷ xs))
+
+------------------------------------------------------------------------
+-- Conversion function
+
+module _ {a} {A : Set a} where
+
+ toList∘fromList : (xs : List A) → toList (fromList xs) ≡ xs
+ toList∘fromList List.[] = refl
+ toList∘fromList (x List.∷ xs) = P.cong (x List.∷_) (toList∘fromList xs)
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+proof-irrelevance-[]= = []=-irrelevance
+{-# WARNING_ON_USAGE proof-irrelevance-[]=
+"Warning: proof-irrelevance-[]= was deprecated in v0.15.
+Please use []=-irrelevance instead."
+#-}
diff --git a/src/Data/Vec/Relation/Equality/DecPropositional.agda b/src/Data/Vec/Relation/Equality/DecPropositional.agda
new file mode 100644
index 0000000..d5b39e6
--- /dev/null
+++ b/src/Data/Vec/Relation/Equality/DecPropositional.agda
@@ -0,0 +1,22 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Decidable vector equality over propositional equality
+------------------------------------------------------------------------
+
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality
+
+module Data.Vec.Relation.Equality.DecPropositional
+ {a} {A : Set a} (_≟_ : Decidable {A = A} _≡_) where
+
+import Data.Vec.Relation.Equality.Propositional as PEq
+import Data.Vec.Relation.Equality.DecSetoid as DSEq
+
+------------------------------------------------------------------------
+-- Publically re-export everything from decSetoid and propositional
+-- equality
+
+open PEq public
+open DSEq (decSetoid _≟_) public
+ using (_≋?_; ≋-isDecEquivalence; ≋-decSetoid)
diff --git a/src/Data/Vec/Relation/Equality/DecSetoid.agda b/src/Data/Vec/Relation/Equality/DecSetoid.agda
new file mode 100644
index 0000000..6fbba08
--- /dev/null
+++ b/src/Data/Vec/Relation/Equality/DecSetoid.agda
@@ -0,0 +1,37 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Decidable semi-heterogeneous vector equality over setoids
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Data.Vec.Relation.Equality.DecSetoid
+ {a ℓ} (DS : DecSetoid a ℓ) where
+
+open import Data.Nat using (ℕ)
+import Data.Vec.Relation.Equality.Setoid as Equality
+import Data.Vec.Relation.Pointwise.Inductive as PW
+open import Level using (_⊔_)
+open import Relation.Binary using (Decidable)
+
+open DecSetoid DS
+
+------------------------------------------------------------------------
+-- Make all definitions from equality available
+
+open Equality setoid public
+
+------------------------------------------------------------------------
+-- Additional properties
+
+infix 4 _≋?_
+
+_≋?_ : ∀ {m n} → Decidable (_≋_ {m} {n})
+_≋?_ = PW.decidable _≟_
+
+≋-isDecEquivalence : ∀ n → IsDecEquivalence (_≋_ {n})
+≋-isDecEquivalence = PW.isDecEquivalence isDecEquivalence
+
+≋-decSetoid : ℕ → DecSetoid a (a ⊔ ℓ)
+≋-decSetoid = PW.decSetoid DS
diff --git a/src/Data/Vec/Relation/Equality/Propositional.agda b/src/Data/Vec/Relation/Equality/Propositional.agda
new file mode 100644
index 0000000..a87df3d
--- /dev/null
+++ b/src/Data/Vec/Relation/Equality/Propositional.agda
@@ -0,0 +1,36 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Vector equality over propositional equality
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Data.Vec.Relation.Equality.Propositional {a} {A : Set a} where
+
+open import Data.Nat.Base using (ℕ; zero; suc; _+_)
+open import Data.Vec
+open import Data.Vec.Relation.Pointwise.Inductive
+ using (Pointwise-≡⇒≡; ≡⇒Pointwise-≡)
+import Data.Vec.Relation.Equality.Setoid as SEq
+open import Relation.Binary.PropositionalEquality
+open import Relation.Binary.HeterogeneousEquality as H using (_≅_)
+
+------------------------------------------------------------------------
+-- Publically re-export everything from setoid equality
+
+open SEq (setoid A) public
+
+------------------------------------------------------------------------
+-- ≋ is propositional
+
+≋⇒≡ : ∀ {n} {xs ys : Vec A n} → xs ≋ ys → xs ≡ ys
+≋⇒≡ = Pointwise-≡⇒≡
+
+≡⇒≋ : ∀ {n} {xs ys : Vec A n} → xs ≡ ys → xs ≋ ys
+≡⇒≋ = ≡⇒Pointwise-≡
+
+≋⇒≅ : ∀ {m n} {xs : Vec A m} {ys : Vec A n} →
+ xs ≋ ys → xs ≅ ys
+≋⇒≅ p with length-equal p
+... | refl = H.≡-to-≅ (≋⇒≡ p)
diff --git a/src/Data/Vec/Relation/Equality/Setoid.agda b/src/Data/Vec/Relation/Equality/Setoid.agda
new file mode 100644
index 0000000..0928a39
--- /dev/null
+++ b/src/Data/Vec/Relation/Equality/Setoid.agda
@@ -0,0 +1,87 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Semi-heterogeneous vector equality over setoids
+------------------------------------------------------------------------
+
+open import Relation.Binary
+
+module Data.Vec.Relation.Equality.Setoid {a ℓ} (S : Setoid a ℓ) where
+
+open import Data.Nat.Base using (ℕ; zero; suc; _+_)
+open import Data.Vec
+open import Data.Vec.Relation.Pointwise.Inductive as PW
+ using (Pointwise)
+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 (_≅_)
+
+open Setoid S renaming (Carrier to A)
+
+------------------------------------------------------------------------
+-- Definition of equality
+
+infix 4 _≋_
+
+_≋_ : ∀ {m n} → REL (Vec A m) (Vec A n) (a ⊔ ℓ)
+_≋_ = Pointwise _≈_
+
+open Pointwise public using ([]; _∷_)
+open PW public using (length-equal)
+
+------------------------------------------------------------------------
+-- Relational properties
+
+≋-refl : ∀ {n} → Reflexive (_≋_ {n})
+≋-refl = PW.refl refl
+
+≋-sym : ∀ {n m} → Sym _≋_ (_≋_ {m} {n})
+≋-sym = PW.sym sym
+
+≋-trans : ∀ {n m o} → Trans (_≋_ {m}) (_≋_ {n} {o}) (_≋_)
+≋-trans = PW.trans trans
+
+≋-isEquivalence : ∀ n → IsEquivalence (_≋_ {n})
+≋-isEquivalence = PW.isEquivalence isEquivalence
+
+≋-setoid : ℕ → Setoid a (a ⊔ ℓ)
+≋-setoid = PW.setoid S
+
+------------------------------------------------------------------------
+-- map
+
+open PW public using ( map⁺)
+
+------------------------------------------------------------------------
+-- ++
+
+open PW public using (++⁺ ; ++⁻ ; ++ˡ⁻; ++ʳ⁻)
+
+++-identityˡ : ∀ {n} (xs : Vec A n) → [] ++ xs ≋ xs
+++-identityˡ _ = ≋-refl
+
+++-identityʳ : ∀ {n} (xs : Vec A n) → xs ++ [] ≋ xs
+++-identityʳ [] = []
+++-identityʳ (x ∷ xs) = refl ∷ ++-identityʳ xs
+
+map-++-commute : ∀ {b m n} {B : Set b}
+ (f : B → A) (xs : Vec B m) {ys : Vec B n} →
+ map f (xs ++ ys) ≋ map f xs ++ map f ys
+map-++-commute f [] = ≋-refl
+map-++-commute f (x ∷ xs) = refl ∷ map-++-commute f xs
+
+------------------------------------------------------------------------
+-- concat
+
+open PW public using (concat⁺; concat⁻)
+
+------------------------------------------------------------------------
+-- replicate
+
+replicate-shiftʳ : ∀ {m} n x (xs : Vec A m) →
+ replicate {n = n} x ++ (x ∷ xs) ≋
+ replicate {n = 1 + n} x ++ xs
+replicate-shiftʳ zero x xs = ≋-refl
+replicate-shiftʳ (suc n) x xs = refl ∷ (replicate-shiftʳ n x xs)
diff --git a/src/Data/Vec/Relation/Pointwise/Extensional.agda b/src/Data/Vec/Relation/Pointwise/Extensional.agda
new file mode 100644
index 0000000..2291358
--- /dev/null
+++ b/src/Data/Vec/Relation/Pointwise/Extensional.agda
@@ -0,0 +1,227 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Extensional pointwise lifting of relations to vectors
+------------------------------------------------------------------------
+
+module Data.Vec.Relation.Pointwise.Extensional where
+
+open import Data.Fin using (zero; suc)
+open import Data.Nat using (zero; suc)
+open import Data.Vec as Vec hiding ([_]; head; tail; map)
+open import Data.Vec.Relation.Pointwise.Inductive as Inductive
+ using ([]; _∷_)
+ renaming (Pointwise to IPointwise)
+open import Level using (_⊔_)
+open import Function using (_∘_)
+open import Function.Equality using (_⟨$⟩_)
+open import Function.Equivalence as Equiv
+ using (_⇔_; ⇔-setoid; equivalence; module Equivalence)
+open import Level using (_⊔_) renaming (zero to ℓ₀)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Relation.Binary.Construct.Closure.Transitive as Plus
+ hiding (equivalent; map)
+open import Relation.Nullary
+import Relation.Nullary.Decidable as Dec
+
+record Pointwise {a b ℓ} {A : Set a} {B : Set b} (_∼_ : REL A B ℓ)
+ {n} (xs : Vec A n) (ys : Vec B n) : Set (a ⊔ b ⊔ ℓ)
+ where
+ constructor ext
+ field app : ∀ i → lookup i xs ∼ lookup i ys
+
+------------------------------------------------------------------------
+-- Operations
+
+head : ∀ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ}
+ {n x y xs} {ys : Vec B n} →
+ Pointwise _∼_ (x ∷ xs) (y ∷ ys) → x ∼ y
+head (ext app) = app zero
+
+tail : ∀ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ}
+ {n x y xs} {ys : Vec B n} →
+ Pointwise _∼_ (x ∷ xs) (y ∷ ys) → Pointwise _∼_ xs ys
+tail (ext app) = ext (app ∘ suc)
+
+map : ∀ {a b ℓ} {A : Set a} {B : Set b} {_~_ _~′_ : REL A B ℓ} {n} →
+ _~_ ⇒ _~′_ → Pointwise _~_ ⇒ Pointwise _~′_ {n}
+map ~⇒~′ xs~ys = ext (~⇒~′ ∘ Pointwise.app xs~ys)
+
+gmap : ∀ {a b ℓ} {A : Set a} {B : Set b}
+ {_~_ : Rel A ℓ} {_~′_ : Rel B ℓ} {f : A → B} {n} →
+ _~_ =[ f ]⇒ _~′_ →
+ Pointwise _~_ =[ Vec.map {n = n} f ]⇒ Pointwise _~′_
+gmap {_} ~⇒~′ {[]} {[]} xs~ys = ext λ()
+gmap {_~′_ = _~′_} ~⇒~′ {x ∷ xs} {y ∷ ys} xs~ys = ext λ
+ { zero → ~⇒~′ (head xs~ys)
+ ; (suc i) → Pointwise.app (gmap {_~′_ = _~′_} ~⇒~′ (tail xs~ys)) i
+ }
+
+------------------------------------------------------------------------
+-- The inductive and extensional definitions are equivalent.
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_~_ : REL A B ℓ} where
+
+ extensional⇒inductive : ∀ {n} {xs : Vec A n} {ys : Vec B n} →
+ Pointwise _~_ xs ys → IPointwise _~_ xs ys
+ extensional⇒inductive {zero} {[]} {[]} xs~ys = []
+ extensional⇒inductive {suc n} {x ∷ xs} {y ∷ ys} xs~ys =
+ (head xs~ys) ∷ extensional⇒inductive (tail xs~ys)
+
+ inductive⇒extensional : ∀ {n} {xs : Vec A n} {ys : Vec B n} →
+ IPointwise _~_ xs ys → Pointwise _~_ xs ys
+ inductive⇒extensional [] = ext λ()
+ inductive⇒extensional (x~y ∷ xs~ys) = ext λ
+ { zero → x~y
+ ; (suc i) → Pointwise.app (inductive⇒extensional xs~ys) i
+ }
+
+ equivalent : ∀ {n} {xs : Vec A n} {ys : Vec B n} →
+ Pointwise _~_ xs ys ⇔ IPointwise _~_ xs ys
+ equivalent = equivalence extensional⇒inductive inductive⇒extensional
+
+------------------------------------------------------------------------
+-- Relational properties
+
+refl : ∀ {a ℓ} {A : Set a} {_~_ : Rel A ℓ} →
+ ∀ {n} → Reflexive _~_ → Reflexive (Pointwise _~_ {n = n})
+refl ~-rfl = ext (λ _ → ~-rfl)
+
+sym : ∀ {a b ℓ} {A : Set a} {B : Set b} {P : REL A B ℓ} {Q : REL B A ℓ}
+ {n} → Sym P Q → Sym (Pointwise P) (Pointwise Q {n = n})
+sym sm xs∼ys = ext λ i → sm (Pointwise.app xs∼ys i)
+
+trans : ∀ {a b c ℓ} {A : Set a} {B : Set b} {C : Set c}
+ {P : REL A B ℓ} {Q : REL B C ℓ} {R : REL A C ℓ} {n} →
+ Trans P Q R →
+ Trans (Pointwise P) (Pointwise Q) (Pointwise R {n = n})
+trans trns xs∼ys ys∼zs = ext λ i →
+ trns (Pointwise.app xs∼ys i) (Pointwise.app ys∼zs i)
+
+decidable : ∀ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ} →
+ Decidable _∼_ → ∀ {n} → Decidable (Pointwise _∼_ {n = n})
+decidable dec xs ys = Dec.map
+ (Setoid.sym (⇔-setoid _) equivalent)
+ (Inductive.decidable dec xs ys)
+
+isEquivalence : ∀ {a ℓ} {A : Set a} {_~_ : Rel A ℓ} →
+ ∀ {n} → IsEquivalence _~_ →
+ IsEquivalence (Pointwise _~_ {n = n})
+isEquivalence equiv = record
+ { refl = refl (IsEquivalence.refl equiv)
+ ; sym = sym (IsEquivalence.sym equiv)
+ ; trans = trans (IsEquivalence.trans equiv)
+ }
+
+isDecEquivalence : ∀ {a ℓ} {A : Set a} {_~_ : Rel A ℓ} →
+ ∀ {n} → IsDecEquivalence _~_ →
+ IsDecEquivalence (Pointwise _~_ {n = n})
+isDecEquivalence decEquiv = record
+ { isEquivalence = isEquivalence (IsDecEquivalence.isEquivalence decEquiv)
+ ; _≟_ = decidable (IsDecEquivalence._≟_ decEquiv)
+ }
+
+------------------------------------------------------------------------
+-- Pointwise _≡_ is equivalent to _≡_.
+
+module _ {a} {A : Set a} where
+
+ Pointwise-≡⇒≡ : ∀ {n} {xs ys : Vec A n} →
+ Pointwise _≡_ xs ys → xs ≡ ys
+ Pointwise-≡⇒≡ {zero} {[]} {[]} (ext app) = P.refl
+ Pointwise-≡⇒≡ {suc n} {x ∷ xs} {y ∷ ys} xs~ys =
+ P.cong₂ _∷_ (head xs~ys) (Pointwise-≡⇒≡ (tail xs~ys))
+
+ ≡⇒Pointwise-≡ : ∀ {n} {xs ys : Vec A n} →
+ xs ≡ ys → Pointwise _≡_ xs ys
+ ≡⇒Pointwise-≡ P.refl = refl P.refl
+
+ Pointwise-≡↔≡ : ∀ {n} {xs ys : Vec A n} →
+ Pointwise _≡_ xs ys ⇔ xs ≡ ys
+ Pointwise-≡↔≡ {ℓ} {A} =
+ Equiv.equivalence Pointwise-≡⇒≡ ≡⇒Pointwise-≡
+
+------------------------------------------------------------------------
+-- Pointwise and Plus commute when the underlying relation is
+-- reflexive.
+module _ {a ℓ} {A : Set a} {_∼_ : Rel A ℓ} where
+
+ ⁺∙⇒∙⁺ : ∀ {n} {xs ys : Vec A n} →
+ Plus (Pointwise _∼_) xs ys → Pointwise (Plus _∼_) xs ys
+ ⁺∙⇒∙⁺ [ ρ≈ρ′ ] = ext (λ x → [ Pointwise.app ρ≈ρ′ x ])
+ ⁺∙⇒∙⁺ (ρ ∼⁺⟨ ρ≈ρ′ ⟩ ρ′≈ρ″) = ext (λ x →
+ _ ∼⁺⟨ Pointwise.app (⁺∙⇒∙⁺ ρ≈ρ′ ) x ⟩
+ Pointwise.app (⁺∙⇒∙⁺ ρ′≈ρ″) x)
+
+ ∙⁺⇒⁺∙ : ∀ {n} {xs ys : Vec A n} → Reflexive _∼_ →
+ Pointwise (Plus _∼_) xs ys → Plus (Pointwise _∼_) xs ys
+ ∙⁺⇒⁺∙ rfl =
+ Plus.map (_⟨$⟩_ (Equivalence.from equivalent)) ∘
+ helper ∘
+ _⟨$⟩_ (Equivalence.to equivalent)
+ where
+ helper : ∀ {n} {xs ys : Vec A n} →
+ IPointwise (Plus _∼_) xs ys → Plus (IPointwise _∼_) xs ys
+ helper [] = [ [] ]
+ helper (_∷_ {x = x} {y = y} {xs = xs} {ys = ys} x∼y xs∼ys) =
+ x ∷ xs ∼⁺⟨ Plus.map (_∷ Inductive.refl rfl) x∼y ⟩
+ y ∷ xs ∼⁺⟨ Plus.map (rfl ∷_) (helper xs∼ys) ⟩∎
+ y ∷ ys ∎
+
+-- ∙⁺⇒⁺∙ cannot be defined if the requirement of reflexivity
+-- is dropped.
+private
+
+ module Counterexample where
+
+ data D : Set where
+ i j x y z : D
+
+ data _R_ : Rel D ℓ₀ where
+ iRj : i R j
+ xRy : x R y
+ yRz : y R z
+
+ xR⁺z : x [ _R_ ]⁺ z
+ xR⁺z =
+ x ∼⁺⟨ [ xRy ] ⟩
+ y ∼⁺⟨ [ yRz ] ⟩∎
+ z ∎
+
+ ix : Vec D 2
+ ix = i ∷ x ∷ []
+
+ jz : Vec D 2
+ jz = j ∷ z ∷ []
+
+ ix∙⁺jz : IPointwise (Plus _R_) ix jz
+ ix∙⁺jz = [ iRj ] ∷ xR⁺z ∷ []
+
+ ¬ix⁺∙jz : ¬ Plus′ (IPointwise _R_) ix jz
+ ¬ix⁺∙jz [ iRj ∷ () ∷ [] ]
+ ¬ix⁺∙jz ((iRj ∷ xRy ∷ []) ∷ [ () ∷ yRz ∷ [] ])
+ ¬ix⁺∙jz ((iRj ∷ xRy ∷ []) ∷ (() ∷ yRz ∷ []) ∷ _)
+
+ counterexample :
+ ¬ (∀ {n} {xs ys : Vec D n} →
+ Pointwise (Plus _R_) xs ys →
+ Plus (Pointwise _R_) xs ys)
+ counterexample ∙⁺⇒⁺∙ =
+ ¬ix⁺∙jz (Equivalence.to Plus.equivalent ⟨$⟩
+ Plus.map (_⟨$⟩_ (Equivalence.to equivalent))
+ (∙⁺⇒⁺∙ (Equivalence.from equivalent ⟨$⟩ ix∙⁺jz)))
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+Pointwise-≡ = Pointwise-≡↔≡
+{-# WARNING_ON_USAGE Pointwise-≡
+"Warning: Pointwise-≡ was deprecated in v0.15.
+Please use Pointwise-≡↔≡ instead."
+#-}
diff --git a/src/Data/Vec/Relation/Pointwise/Inductive.agda b/src/Data/Vec/Relation/Pointwise/Inductive.agda
new file mode 100644
index 0000000..d837c09
--- /dev/null
+++ b/src/Data/Vec/Relation/Pointwise/Inductive.agda
@@ -0,0 +1,258 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Inductive pointwise lifting of relations to vectors
+------------------------------------------------------------------------
+
+module Data.Vec.Relation.Pointwise.Inductive where
+
+open import Algebra.FunctionProperties
+open import Data.Fin using (Fin; zero; suc)
+open import Data.Nat using (ℕ; zero; suc)
+open import Data.Product using (_×_; _,_)
+open import Data.Vec as Vec hiding ([_]; head; tail; map; lookup)
+open import Data.Vec.All using (All; []; _∷_)
+open import Level using (_⊔_)
+open import Function using (_∘_)
+open import Function.Equivalence using (_⇔_; equivalence)
+open import Relation.Binary
+open import Relation.Binary.PropositionalEquality as P using (_≡_)
+open import Relation.Nullary
+
+------------------------------------------------------------------------
+-- Relation
+
+infixr 5 _∷_
+
+data Pointwise {a b ℓ} {A : Set a} {B : Set b} (_∼_ : REL A B ℓ) :
+ ∀ {m n} (xs : Vec A m) (ys : Vec B n) → Set (a ⊔ b ⊔ ℓ)
+ where
+ [] : Pointwise _∼_ [] []
+ _∷_ : ∀ {m n x y} {xs : Vec A m} {ys : Vec B n}
+ (x∼y : x ∼ y) (xs∼ys : Pointwise _∼_ xs ys) →
+ Pointwise _∼_ (x ∷ xs) (y ∷ ys)
+
+length-equal : ∀ {a b m n ℓ} {A : Set a} {B : Set b} {_~_ : REL A B ℓ}
+ {xs : Vec A m} {ys : Vec B n} →
+ Pointwise _~_ xs ys → m ≡ n
+length-equal [] = P.refl
+length-equal (_ ∷ xs~ys) = P.cong suc (length-equal xs~ys)
+
+------------------------------------------------------------------------
+-- Operations
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_~_ : REL A B ℓ} where
+
+ head : ∀ {m n x y} {xs : Vec A m} {ys : Vec B n} →
+ Pointwise _~_ (x ∷ xs) (y ∷ ys) → x ~ y
+ head (x∼y ∷ xs∼ys) = x∼y
+
+ tail : ∀ {m n x y} {xs : Vec A m} {ys : Vec B n} →
+ Pointwise _~_ (x ∷ xs) (y ∷ ys) → Pointwise _~_ xs ys
+ tail (x∼y ∷ xs∼ys) = xs∼ys
+
+ lookup : ∀ {n} {xs : Vec A n} {ys : Vec B n} → Pointwise _~_ xs ys →
+ ∀ i → (Vec.lookup i xs) ~ (Vec.lookup i ys)
+ lookup (x~y ∷ _) zero = x~y
+ lookup (_ ∷ xs~ys) (suc i) = lookup xs~ys i
+
+ map : ∀ {ℓ₂} {_≈_ : REL A B ℓ₂} →
+ _≈_ ⇒ _~_ → ∀ {m n} → Pointwise _≈_ ⇒ Pointwise _~_ {m} {n}
+ map ~₁⇒~₂ [] = []
+ map ~₁⇒~₂ (x∼y ∷ xs~ys) = ~₁⇒~₂ x∼y ∷ map ~₁⇒~₂ xs~ys
+
+------------------------------------------------------------------------
+-- Relational properties
+
+refl : ∀ {a ℓ} {A : Set a} {_~_ : Rel A ℓ} →
+ ∀ {n} → Reflexive _~_ → Reflexive (Pointwise _~_ {n})
+refl ~-refl {[]} = []
+refl ~-refl {x ∷ xs} = ~-refl ∷ refl ~-refl
+
+sym : ∀ {a b ℓ} {A : Set a} {B : Set b}
+ {P : REL A B ℓ} {Q : REL B A ℓ} {m n} →
+ Sym P Q → Sym (Pointwise P) (Pointwise Q {m} {n})
+sym sm [] = []
+sym sm (x∼y ∷ xs~ys) = sm x∼y ∷ sym sm xs~ys
+
+trans : ∀ {a b c ℓ} {A : Set a} {B : Set b} {C : Set c}
+ {P : REL A B ℓ} {Q : REL B C ℓ} {R : REL A C ℓ} {m n o} →
+ Trans P Q R →
+ Trans (Pointwise P {m}) (Pointwise Q {n} {o}) (Pointwise R)
+trans trns [] [] = []
+trans trns (x∼y ∷ xs~ys) (y∼z ∷ ys~zs) =
+ trns x∼y y∼z ∷ trans trns xs~ys ys~zs
+
+decidable : ∀ {a b ℓ} {A : Set a} {B : Set b} {_∼_ : REL A B ℓ} →
+ Decidable _∼_ → ∀ {m n} → Decidable (Pointwise _∼_ {m} {n})
+decidable dec [] [] = yes []
+decidable dec [] (y ∷ ys) = no λ()
+decidable dec (x ∷ xs) [] = no λ()
+decidable dec (x ∷ xs) (y ∷ ys) with dec x y
+... | no ¬x∼y = no (¬x∼y ∘ head)
+... | yes x∼y with decidable dec xs ys
+... | no ¬xs∼ys = no (¬xs∼ys ∘ tail)
+... | yes xs∼ys = yes (x∼y ∷ xs∼ys)
+
+isEquivalence : ∀ {a ℓ} {A : Set a} {_~_ : Rel A ℓ} →
+ IsEquivalence _~_ → ∀ n →
+ IsEquivalence (Pointwise _~_ {n})
+isEquivalence equiv n = record
+ { refl = refl (IsEquivalence.refl equiv)
+ ; sym = sym (IsEquivalence.sym equiv)
+ ; trans = trans (IsEquivalence.trans equiv)
+ }
+
+isDecEquivalence : ∀ {a ℓ} {A : Set a} {_~_ : Rel A ℓ} →
+ IsDecEquivalence _~_ → ∀ n →
+ IsDecEquivalence (Pointwise _~_ {n})
+isDecEquivalence decEquiv n = record
+ { isEquivalence = isEquivalence (IsDecEquivalence.isEquivalence decEquiv) n
+ ; _≟_ = decidable (IsDecEquivalence._≟_ decEquiv)
+ }
+
+setoid : ∀ {a ℓ} → Setoid a ℓ → ℕ → Setoid a (a ⊔ ℓ)
+setoid S n = record
+ { isEquivalence = isEquivalence (Setoid.isEquivalence S) n
+ }
+
+decSetoid : ∀ {a ℓ} → DecSetoid a ℓ → ℕ → DecSetoid a (a ⊔ ℓ)
+decSetoid S n = record
+ { isDecEquivalence = isDecEquivalence (DecSetoid.isDecEquivalence S) n
+ }
+
+------------------------------------------------------------------------
+-- map
+
+module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} where
+
+ map⁺ : ∀ {ℓ₁ ℓ₂} {_~₁_ : REL A B ℓ₁} {_~₂_ : REL C D ℓ₂}
+ {f : A → C} {g : B → D} →
+ (∀ {x y} → x ~₁ y → f x ~₂ g y) →
+ ∀ {m n xs ys} → Pointwise _~₁_ {m} {n} xs ys →
+ Pointwise _~₂_ (Vec.map f xs) (Vec.map g ys)
+ map⁺ ~₁⇒~₂ [] = []
+ map⁺ ~₁⇒~₂ (x∼y ∷ xs~ys) = ~₁⇒~₂ x∼y ∷ map⁺ ~₁⇒~₂ xs~ys
+
+------------------------------------------------------------------------
+-- _++_
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_~_ : REL A B ℓ} where
+
+ ++⁺ : ∀ {m n p q}
+ {ws : Vec A m} {xs : Vec B p} {ys : Vec A n} {zs : Vec B q} →
+ Pointwise _~_ ws xs → Pointwise _~_ ys zs →
+ Pointwise _~_ (ws ++ ys) (xs ++ zs)
+ ++⁺ [] ys~zs = ys~zs
+ ++⁺ (w~x ∷ ws~xs) ys~zs = w~x ∷ (++⁺ ws~xs ys~zs)
+
+ ++ˡ⁻ : ∀ {m n}
+ (ws : Vec A m) (xs : Vec B m) {ys : Vec A n} {zs : Vec B n} →
+ Pointwise _~_ (ws ++ ys) (xs ++ zs) → Pointwise _~_ ws xs
+ ++ˡ⁻ [] [] _ = []
+ ++ˡ⁻ (w ∷ ws) (x ∷ xs) (w~x ∷ ps) = w~x ∷ ++ˡ⁻ ws xs ps
+
+ ++ʳ⁻ : ∀ {m n}
+ (ws : Vec A m) (xs : Vec B m) {ys : Vec A n} {zs : Vec B n} →
+ Pointwise _~_ (ws ++ ys) (xs ++ zs) → Pointwise _~_ ys zs
+ ++ʳ⁻ [] [] ys~zs = ys~zs
+ ++ʳ⁻ (w ∷ ws) (x ∷ xs) (_ ∷ ps) = ++ʳ⁻ ws xs ps
+
+ ++⁻ : ∀ {m n}
+ (ws : Vec A m) (xs : Vec B m) {ys : Vec A n} {zs : Vec B n} →
+ Pointwise _~_ (ws ++ ys) (xs ++ zs) →
+ Pointwise _~_ ws xs × Pointwise _~_ ys zs
+ ++⁻ ws xs ps = ++ˡ⁻ ws xs ps , ++ʳ⁻ ws xs ps
+
+------------------------------------------------------------------------
+-- concat
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_~_ : REL A B ℓ} where
+
+ concat⁺ : ∀ {m n p q}
+ {xss : Vec (Vec A m) n} {yss : Vec (Vec B p) q} →
+ Pointwise (Pointwise _~_) xss yss →
+ Pointwise _~_ (concat xss) (concat yss)
+ concat⁺ [] = []
+ concat⁺ (xs~ys ∷ ps) = ++⁺ xs~ys (concat⁺ ps)
+
+ concat⁻ : ∀ {m n} (xss : Vec (Vec A m) n) (yss : Vec (Vec B m) n) →
+ Pointwise _~_ (concat xss) (concat yss) →
+ Pointwise (Pointwise _~_) xss yss
+ concat⁻ [] [] [] = []
+ concat⁻ (xs ∷ xss) (ys ∷ yss) ps =
+ ++ˡ⁻ xs ys ps ∷ concat⁻ xss yss (++ʳ⁻ xs ys ps)
+
+------------------------------------------------------------------------
+-- tabulate
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {_~_ : REL A B ℓ} where
+
+ tabulate⁺ : ∀ {n} {f : Fin n → A} {g : Fin n → B} →
+ (∀ i → f i ~ g i) →
+ Pointwise _~_ (tabulate f) (tabulate g)
+ tabulate⁺ {zero} f~g = []
+ tabulate⁺ {suc n} f~g = f~g zero ∷ tabulate⁺ (f~g ∘ suc)
+
+ tabulate⁻ : ∀ {n} {f : Fin n → A} {g : Fin n → B} →
+ Pointwise _~_ (tabulate f) (tabulate g) →
+ (∀ i → f i ~ g i)
+ tabulate⁻ (f₀~g₀ ∷ _) zero = f₀~g₀
+ tabulate⁻ (_ ∷ f~g) (suc i) = tabulate⁻ f~g i
+
+------------------------------------------------------------------------
+-- Degenerate pointwise relations
+
+module _ {a b ℓ} {A : Set a} {B : Set b} {P : A → Set ℓ} where
+
+ Pointwiseˡ⇒All : ∀ {m n} {xs : Vec A m} {ys : Vec B n} →
+ Pointwise (λ x y → P x) xs ys → All P xs
+ Pointwiseˡ⇒All [] = []
+ Pointwiseˡ⇒All (p ∷ ps) = p ∷ Pointwiseˡ⇒All ps
+
+ Pointwiseʳ⇒All : ∀ {n} {xs : Vec B n} {ys : Vec A n} →
+ Pointwise (λ x y → P y) xs ys → All P ys
+ Pointwiseʳ⇒All [] = []
+ Pointwiseʳ⇒All (p ∷ ps) = p ∷ Pointwiseʳ⇒All ps
+
+ All⇒Pointwiseˡ : ∀ {n} {xs : Vec A n} {ys : Vec B n} →
+ All P xs → Pointwise (λ x y → P x) xs ys
+ All⇒Pointwiseˡ {ys = []} [] = []
+ All⇒Pointwiseˡ {ys = _ ∷ _} (p ∷ ps) = p ∷ All⇒Pointwiseˡ ps
+
+ All⇒Pointwiseʳ : ∀ {n} {xs : Vec B n} {ys : Vec A n} →
+ All P ys → Pointwise (λ x y → P y) xs ys
+ All⇒Pointwiseʳ {xs = []} [] = []
+ All⇒Pointwiseʳ {xs = _ ∷ _} (p ∷ ps) = p ∷ All⇒Pointwiseʳ ps
+
+------------------------------------------------------------------------
+-- Pointwise _≡_ is equivalent to _≡_
+
+module _ {a} {A : Set a} where
+
+ Pointwise-≡⇒≡ : ∀ {n} {xs ys : Vec A n} →
+ Pointwise _≡_ xs ys → xs ≡ ys
+ Pointwise-≡⇒≡ [] = P.refl
+ Pointwise-≡⇒≡ (P.refl ∷ xs~ys) = P.cong (_ ∷_) (Pointwise-≡⇒≡ xs~ys)
+
+ ≡⇒Pointwise-≡ : ∀ {n} {xs ys : Vec A n} →
+ xs ≡ ys → Pointwise _≡_ xs ys
+ ≡⇒Pointwise-≡ P.refl = refl P.refl
+
+ Pointwise-≡↔≡ : ∀ {n} {xs ys : Vec A n} →
+ Pointwise _≡_ xs ys ⇔ xs ≡ ys
+ Pointwise-≡↔≡ = equivalence Pointwise-≡⇒≡ ≡⇒Pointwise-≡
+
+------------------------------------------------------------------------
+-- DEPRECATED NAMES
+------------------------------------------------------------------------
+-- Please use the new names as continuing support for the old names is
+-- not guaranteed.
+
+-- Version 0.15
+
+Pointwise-≡ = Pointwise-≡↔≡
+{-# WARNING_ON_USAGE Pointwise-≡
+"Warning: Pointwise-≡ was deprecated in v0.15.
+Please use Pointwise-≡↔≡ instead."
+#-}
diff --git a/src/Data/W.agda b/src/Data/W.agda
index 4d47f7b..d0128aa 100644
--- a/src/Data/W.agda
+++ b/src/Data/W.agda
@@ -7,25 +7,61 @@
module Data.W where
open import Level
+open import Function
+open import Data.Product hiding (map)
+open import Data.Container.Core
open import Relation.Nullary
+open import Agda.Builtin.Equality
-- The family of W-types.
-data W {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
- sup : (x : A) (f : B x → W A B) → W A B
+data W {s p} (C : Container s p) : Set (s ⊔ p) where
+ sup : ⟦ C ⟧ (W C) → W C
+
+module _ {s p} {C : Container s p} (open Container C)
+ {s : Shape} {f : Position s → W C} where
+
+ sup-injective₁ : ∀ {t g} → sup (s , f) ≡ sup (t , g) → s ≡ t
+ sup-injective₁ refl = refl
+
+ sup-injective₂ : ∀ {g} → sup (s , f) ≡ sup (s , g) → f ≡ g
+ sup-injective₂ refl = refl
-- Projections.
-head : ∀ {a b} {A : Set a} {B : A → Set b} →
- W A B → A
-head (sup x f) = x
+module _ {s p} {C : Container s p} (open Container C) where
+
+ head : W C → Shape
+ head (sup (x , f)) = x
+
+ tail : (x : W C) → Position (head x) → W C
+ tail (sup (x , f)) = f
+
+-- map
+
+module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂}
+ (m : C₁ ⇒ C₂) where
+
+ map : W C₁ → W C₂
+ map (sup (x , f)) = sup (⟪ m ⟫ (x , λ p → map (f p)))
+
+-- induction
+
+module _ {s p ℓ} {C : Container s p} (P : W C → Set ℓ)
+ (alg : ∀ {t} → □ P t → P (sup t)) where
+
+ induction : (w : W C) → P w
+ induction (sup (s , f)) = alg $ λ p → induction (f p)
+
+module _ {s p ℓ} {C : Container s p} (open Container C)
+ {P : Set ℓ} (alg : ⟦ C ⟧ P → P) where
+
+ foldr : W C → P
+ foldr = induction (const P) (λ p → alg (_ , p))
-tail : ∀ {a b} {A : Set a} {B : A → Set b} →
- (x : W A B) → B (head x) → W A B
-tail (sup x f) = f
+-- If Position is always inhabited, then W_C is empty.
--- If B is always inhabited, then W A B is empty.
+module _ {s p} {C : Container s p} (open Container C) where
-inhabited⇒empty : ∀ {a b} {A : Set a} {B : A → Set b} →
- (∀ x → B x) → ¬ W A B
-inhabited⇒empty b (sup x f) = inhabited⇒empty b (f (b x))
+ inhabited⇒empty : (∀ s → Position s) → ¬ W C
+ inhabited⇒empty b = foldr ((_$ b _) ∘ proj₂)
diff --git a/src/Data/Char/Core.agda b/src/Data/Word.agda
index a31c435..ea5f0a9 100644
--- a/src/Data/Char/Core.agda
+++ b/src/Data/Word.agda
@@ -1,12 +1,17 @@
------------------------------------------------------------------------
-- The Agda standard library
--
--- Basic definitions for Characters
+-- Machine words
------------------------------------------------------------------------
-module Data.Char.Core where
+module Data.Word where
------------------------------------------------------------------------
--- The type
+-- Re-export built-ins publically
-open import Agda.Builtin.Char public using (Char)
+open import Agda.Builtin.Word public
+ using (Word64)
+ renaming
+ ( primWord64ToNat to toℕ
+ ; primWord64FromNat to fromℕ
+ )
diff --git a/src/Data/Word/Unsafe.agda b/src/Data/Word/Unsafe.agda
new file mode 100644
index 0000000..b57f996
--- /dev/null
+++ b/src/Data/Word/Unsafe.agda
@@ -0,0 +1,22 @@
+------------------------------------------------------------------------
+-- The Agda standard library
+--
+-- Unsafe machine word operations
+------------------------------------------------------------------------
+
+module Data.Word.Unsafe where
+
+import Data.Nat as ℕ
+open import Data.Word using (Word64; toℕ)
+open import Relation.Nullary using (Dec; yes; no)
+open import Relation.Binary.PropositionalEquality using (_≡_; refl)
+open import Relation.Binary.PropositionalEquality.TrustMe
+
+------------------------------------------------------------------------
+-- An informative equality test.
+
+_≟_ : (a b : Word64) → Dec (a ≡ b)
+a ≟ b with toℕ a ℕ.≟ toℕ b
+... | yes _ = yes trustMe
+... | no _ = no whatever
+ where postulate whatever : _