diff --git a/CHANGELOG.md b/CHANGELOG.md index 907f8c5be2..a39702af58 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -16,6 +16,14 @@ Bug-fixes Non-backwards compatible changes -------------------------------- +* The modules and morphisms in `Algebra.Module.Morphism.Structures` are now + parametrized by _raw_ bundles rather than lawful bundles, in line with other + modules defining morphism structures. +* The definitions in `Algebra.Module.Morphism.Construct.Composition` are now + parametrized by _raw_ bundles, and as such take a proof of transitivity. +* The definitions in `Algebra.Module.Morphism.Construct.Identity` are now + parametrized by _raw_ bundles, and as such take a proof of reflexivity. + Other major improvements ------------------------ @@ -25,6 +33,11 @@ Deprecated modules Deprecated names ---------------- +* In `Algebra.Properties.Semiring.Mult`: + ```agda + 1×-identityʳ ↦ ×-homo-1 + ``` + * In `Data.Nat.Divisibility.Core`: ```agda *-pres-∣ ↦ Data.Nat.Divisibility.*-pres-∣ @@ -33,6 +46,8 @@ Deprecated names New modules ----------- +* `Algebra.Module.Bundles.Raw`: raw bundles for module-like algebraic structures + * Integer arithmetic modulo `n`, based on `Data.Nat.Bounded.*`: ```agda Data.Integer.Modulo.Base @@ -43,6 +58,64 @@ New modules Additions to existing modules ----------------------------- +* Exporting more `Raw` substructures from `Algebra.Bundles.Ring`: + ```agda + rawNearSemiring : RawNearSemiring _ _ + rawRingWithoutOne : RawRingWithoutOne _ _ + +-rawGroup : RawGroup _ _ + ``` + +* In `Algebra.Module.Bundles`, raw bundles are re-exported and the bundles expose their raw counterparts. + +* In `Algebra.Module.Construct.DirectProduct`: + ```agda + rawLeftSemimodule : RawLeftSemimodule R m ℓm → RawLeftSemimodule m′ ℓm′ → RawLeftSemimodule R (m ⊔ m′) (ℓm ⊔ ℓm′) + rawLeftModule : RawLeftModule R m ℓm → RawLeftModule m′ ℓm′ → RawLeftModule R (m ⊔ m′) (ℓm ⊔ ℓm′) + rawRightSemimodule : RawRightSemimodule R m ℓm → RawRightSemimodule m′ ℓm′ → RawRightSemimodule R (m ⊔ m′) (ℓm ⊔ ℓm′) + rawRightModule : RawRightModule R m ℓm → RawRightModule m′ ℓm′ → RawRightModule R (m ⊔ m′) (ℓm ⊔ ℓm′) + rawBisemimodule : RawBisemimodule R m ℓm → RawBisemimodule m′ ℓm′ → RawBisemimodule R (m ⊔ m′) (ℓm ⊔ ℓm′) + rawBimodule : RawBimodule R m ℓm → RawBimodule m′ ℓm′ → RawBimodule R (m ⊔ m′) (ℓm ⊔ ℓm′) + rawSemimodule : RawSemimodule R m ℓm → RawSemimodule m′ ℓm′ → RawSemimodule R (m ⊔ m′) (ℓm ⊔ ℓm′) + rawModule : RawModule R m ℓm → RawModule m′ ℓm′ → RawModule R (m ⊔ m′) (ℓm ⊔ ℓm′) + ``` + +* In `Algebra.Module.Construct.TensorUnit`: + ```agda + rawLeftSemimodule : RawLeftSemimodule _ c ℓ + rawLeftModule : RawLeftModule _ c ℓ + rawRightSemimodule : RawRightSemimodule _ c ℓ + rawRightModule : RawRightModule _ c ℓ + rawBisemimodule : RawBisemimodule _ _ c ℓ + rawBimodule : RawBimodule _ _ c ℓ + rawSemimodule : RawSemimodule _ c ℓ + rawModule : RawModule _ c ℓ + ``` + +* In `Algebra.Module.Construct.Zero`: + ```agda + rawLeftSemimodule : RawLeftSemimodule R c ℓ + rawLeftModule : RawLeftModule R c ℓ + rawRightSemimodule : RawRightSemimodule R c ℓ + rawRightModule : RawRightModule R c ℓ + rawBisemimodule : RawBisemimodule R c ℓ + rawBimodule : RawBimodule R c ℓ + rawSemimodule : RawSemimodule R c ℓ + rawModule : RawModule R c ℓ + ``` + +* In `Algebra.Properties.Monoid.Mult`: + ```agda + ×-homo-0 : ∀ x → 0 × x ≈ 0# + ×-homo-1 : ∀ x → 1 × x ≈ x + ``` + +* In `Algebra.Properties.Semiring.Mult`: + ```agda + ×-homo-0# : ∀ x → 0 × x ≈ 0# * x + ×-homo-1# : ∀ x → 1 × x ≈ 1# * x + idem-×-homo-* : (_*_ IdempotentOn x) → (m × x) * (n × x) ≈ (m ℕ.* n) × x + ``` + * In `Data.Fin.Properties`: ```agda nonZeroIndex : Fin n → ℕ.NonZero n @@ -99,3 +172,25 @@ Additions to existing modules * In `Function.Bundles`, added `_⟶ₛ_` as a synonym for `Func` that can be used infix. + +* Added new definitions in `Relation.Binary` + ``` + Stable : Pred A ℓ → Set _ + ``` + +* Added new definitions in `Relation.Nullary` + ``` + Recomputable : Set _ + WeaklyDecidable : Set _ + ``` + +* Added new definitions in `Relation.Unary` + ``` + Stable : Pred A ℓ → Set _ + WeaklyDecidable : Pred A ℓ → Set _ + ``` + +* Added new proof in `Relation.Nullary.Decidable`: + ```agda + ⌊⌋-map′ : (a? : Dec A) → ⌊ map′ t f a? ⌋ ≡ ⌊ a? ⌋ + ``` diff --git a/doc/README.agda b/doc/README.agda index 41b536a0d0..6a9a3e398c 100644 --- a/doc/README.agda +++ b/doc/README.agda @@ -222,6 +222,10 @@ import Text.Printf import README.Case +-- Showcasing the framework for well-scoped substitutions + +import README.Data.Fin.Substitution.UntypedLambda + -- Some examples showing how combinators can be used to emulate -- "functional reasoning" diff --git a/doc/README/Data/Fin/Substitution/UntypedLambda.agda b/doc/README/Data/Fin/Substitution/UntypedLambda.agda index 126959958f..036f7ed5ca 100644 --- a/doc/README/Data/Fin/Substitution/UntypedLambda.agda +++ b/doc/README/Data/Fin/Substitution/UntypedLambda.agda @@ -14,7 +14,7 @@ open import Data.Fin.Substitution.Lemmas open import Data.Nat.Base hiding (_/_) open import Data.Fin.Base using (Fin) open import Data.Vec.Base -open import Relation.Binary.PropositionalEquality.Core +open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; cong; cong₂; module ≡-Reasoning) open import Relation.Binary.Construct.Closure.ReflexiveTransitive using (Star; ε; _◅_) diff --git a/doc/style-guide.md b/doc/style-guide.md index 5732747cda..0e3a54841a 100644 --- a/doc/style-guide.md +++ b/doc/style-guide.md @@ -121,8 +121,9 @@ automate most of this. * If it is important that certain names only come into scope later in the file then the module should still be imported at the top of the - file but it can be given a shorter name using the keyword `as` and then - opened later on in the file when needed, e.g. + file but it can be imported *qualified*, i.e. given a shorter name + using the keyword `as` and then opened later on in the file when needed, + e.g. ```agda import Data.List.Relation.Binary.Equality.Setoid as SetoidEquality ... @@ -130,6 +131,21 @@ automate most of this. open SetoidEquality S ``` +* Naming conventions for qualified `import`s: if importing a module under + a root of the form `Data.X` (e.g. the `Base` module for basic operations, + or `Properties` for lemmas about them etc.) then conventionally, the + qualified name(s) for the import(s) should (all) share as qualfied name + that of the name of the `X` datatype defined: i.e. `Data.Nat.Base` + should be imported as `ℕ`, `Data.List.Properties` as `List`, etc. + In this spirit, the convention applies also to (the datatype defined by) + `Relation.Binary.PropositionalEquality.*` which should be imported qualified + with the name `≡`. + Other modules should be given a 'suitable' qualified name based on its 'long' + path-derived name (such as `SetoidEquality` in the example above); commonly + occcurring examples such as `Algebra.Structures` should be imported qualified + as `Structures` etc. + NB. Historical legacy means that these conventions have not always been observed! + * When using only a few items (i.e. < 5) from a module, it is a good practice to enumerate the items that will be used by declaring the import statement with the directive `using`. This makes the dependencies clearer, e.g. @@ -532,3 +548,22 @@ word within a compound word is capitalized except for the first word. * The names of patterns for reflected syntax are also *appended* with an additional backtick. + +#### Specific pragmatics/idiomatic patterns + +## Use of `with` notation + +Thinking on this has changed since the early days of the library, with +a desire to avoid 'unnecessary' uses of `with`: see Issues +[#1937](https://github.com/agda/agda-stdlib/issues/1937) and +[#2123](https://github.com/agda/agda-stdlib/issues/2123). + +## Proving instances of `Decidable` for sets, predicates, relations, ... + +Issue [#803](https://github.com/agda/agda-stdlib/issues/803) +articulates a programming pattern for writing proofs of decidability, +used successfully in PR +[#799](https://github.com/agda/agda-stdlib/pull/799) and made +systematic for `Nary` relations in PR +[#811](https://github.com/agda/agda-stdlib/pull/811) + diff --git a/src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda b/src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda index 868224bf0c..d8b3ffcdfd 100644 --- a/src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda +++ b/src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda @@ -21,7 +21,7 @@ open CommutativeRing commutativeRing using (ring; *-commutativeMonoid) open import Algebra.Properties.Ring ring using (-0#≈0#; -‿distribˡ-*; -‿distribʳ-*; -‿anti-homo-+; -‿involutive) open import Relation.Binary.Definitions using (Symmetric) -import Relation.Binary.Reasoning.Setoid as ReasonSetoid +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning open import Algebra.Properties.CommutativeMonoid *-commutativeMonoid private variable @@ -50,7 +50,7 @@ x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0) helper (x⁻¹ , x⁻¹*x≈1 , x*x⁻¹≈1) (y⁻¹ , y⁻¹*y≈1 , y*y⁻¹≈1) = invertibleˡ⇒# (y⁻¹ * x⁻¹ , y⁻¹*x⁻¹*x*y≈1) where - open ReasonSetoid setoid + open ≈-Reasoning setoid y⁻¹*x⁻¹*x*y≈1 : y⁻¹ * x⁻¹ * (x * y - 0#) ≈ 1# y⁻¹*x⁻¹*x*y≈1 = begin @@ -67,7 +67,7 @@ x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0) #-sym : Symmetric _#_ #-sym {x} {y} x#y = invertibleˡ⇒# (- x-y⁻¹ , x-y⁻¹*y-x≈1) where - open ReasonSetoid setoid + open ≈-Reasoning setoid InvX-Y : Invertible _≈_ 1# _*_ (x - y) InvX-Y = #⇒invertible x#y @@ -96,7 +96,7 @@ x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0) helper (x-z⁻¹ , x-z⁻¹*x-z≈1# , x-z*x-z⁻¹≈1#) = invertibleˡ⇒# (x-z⁻¹ , x-z⁻¹*y-z≈1) where - open ReasonSetoid setoid + open ≈-Reasoning setoid x-z⁻¹*y-z≈1 : x-z⁻¹ * (y - z) ≈ 1# x-z⁻¹*y-z≈1 = begin diff --git a/src/Algebra/Bundles.agda b/src/Algebra/Bundles.agda index 6b6384458a..1a31d16496 100644 --- a/src/Algebra/Bundles.agda +++ b/src/Algebra/Bundles.agda @@ -961,6 +961,9 @@ record Ring c ℓ : Set (suc (c ⊔ ℓ)) where ; semiringWithoutAnnihilatingZero ) + open NearSemiring nearSemiring public + using (rawNearSemiring) + open AbelianGroup +-abelianGroup public using () renaming (group to +-group; invertibleMagma to +-invertibleMagma; invertibleUnitalMagma to +-invertibleUnitalMagma) @@ -974,6 +977,9 @@ record Ring c ℓ : Set (suc (c ⊔ ℓ)) where ; 1# = 1# } + open RawRing rawRing public + using (rawRingWithoutOne; +-rawGroup) + record CommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where infix 8 -_ diff --git a/src/Algebra/Construct/Flip/Op.agda b/src/Algebra/Construct/Flip/Op.agda index 4dd5849a6f..dd139e4ba1 100644 --- a/src/Algebra/Construct/Flip/Op.agda +++ b/src/Algebra/Construct/Flip/Op.agda @@ -10,7 +10,7 @@ module Algebra.Construct.Flip.Op where open import Algebra -import Data.Product.Base as Prod +import Data.Product.Base as Product import Data.Sum.Base as Sum open import Function.Base using (flip) open import Level using (Level) @@ -39,7 +39,7 @@ module _ (≈ : Rel A ℓ) (∙ : Op₂ A) where associative sym assoc x y z = sym (assoc z y x) identity : Identity ≈ ε ∙ → Identity ≈ ε (flip ∙) - identity id = Prod.swap id + identity id = Product.swap id commutative : Commutative ≈ ∙ → Commutative ≈ (flip ∙) commutative comm = flip comm @@ -51,7 +51,7 @@ module _ (≈ : Rel A ℓ) (∙ : Op₂ A) where idempotent idem = idem inverse : Inverse ≈ ε ⁻¹ ∙ → Inverse ≈ ε ⁻¹ (flip ∙) - inverse inv = Prod.swap inv + inverse inv = Product.swap inv ------------------------------------------------------------------------ -- Structures diff --git a/src/Algebra/Construct/LexProduct.agda b/src/Algebra/Construct/LexProduct.agda index e4d51ff9ed..e2f5825824 100644 --- a/src/Algebra/Construct/LexProduct.agda +++ b/src/Algebra/Construct/LexProduct.agda @@ -16,7 +16,7 @@ open import Relation.Binary.Core using (Rel) open import Relation.Binary.Definitions using (Decidable) open import Relation.Nullary using (¬_; does; yes; no) open import Relation.Nullary.Negation using (contradiction; contradiction₂) -import Relation.Binary.Reasoning.Setoid as SetoidReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning module Algebra.Construct.LexProduct {ℓ₁ ℓ₂ ℓ₃ ℓ₄} (M : Magma ℓ₁ ℓ₂) (N : Magma ℓ₃ ℓ₄) diff --git a/src/Algebra/Construct/LiftedChoice.agda b/src/Algebra/Construct/LiftedChoice.agda index f53eecc7f1..9b141f8f7e 100644 --- a/src/Algebra/Construct/LiftedChoice.agda +++ b/src/Algebra/Construct/LiftedChoice.agda @@ -20,7 +20,7 @@ open import Level using (Level; _⊔_) open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) open import Relation.Unary using (Pred) -import Relation.Binary.Reasoning.Setoid as EqReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning private variable @@ -46,7 +46,7 @@ module _ {_≈_ : Rel B ℓ} {_∙_ : Op₂ B} private module M = IsSelectiveMagma ∙-isSelectiveMagma open M hiding (sel; isMagma) - open EqReasoning setoid + open ≈-Reasoning setoid module _ (f : A → B) where diff --git a/src/Algebra/Construct/Subst/Equality.agda b/src/Algebra/Construct/Subst/Equality.agda index 05d0277398..6dc1ced9e2 100644 --- a/src/Algebra/Construct/Subst/Equality.agda +++ b/src/Algebra/Construct/Subst/Equality.agda @@ -9,7 +9,7 @@ {-# OPTIONS --cubical-compatible --safe #-} -open import Data.Product.Base as Prod +open import Data.Product.Base as Product open import Relation.Binary.Core module Algebra.Construct.Subst.Equality @@ -45,13 +45,13 @@ sel : ∀ {∙} → Selective ≈₁ ∙ → Selective ≈₂ ∙ sel sel x y = Sum.map to to (sel x y) identity : ∀ {∙ e} → Identity ≈₁ e ∙ → Identity ≈₂ e ∙ -identity = Prod.map (to ∘_) (to ∘_) +identity = Product.map (to ∘_) (to ∘_) inverse : ∀ {∙ e ⁻¹} → Inverse ≈₁ ⁻¹ ∙ e → Inverse ≈₂ ⁻¹ ∙ e -inverse = Prod.map (to ∘_) (to ∘_) +inverse = Product.map (to ∘_) (to ∘_) absorptive : ∀ {∙ ◦} → Absorptive ≈₁ ∙ ◦ → Absorptive ≈₂ ∙ ◦ -absorptive = Prod.map (λ f x y → to (f x y)) (λ f x y → to (f x y)) +absorptive = Product.map (λ f x y → to (f x y)) (λ f x y → to (f x y)) distribˡ : ∀ {∙ ◦} → _DistributesOverˡ_ ≈₁ ∙ ◦ → _DistributesOverˡ_ ≈₂ ∙ ◦ distribˡ distribˡ x y z = to (distribˡ x y z) @@ -60,7 +60,7 @@ distribʳ : ∀ {∙ ◦} → _DistributesOverʳ_ ≈₁ ∙ ◦ → _Distribute distribʳ distribʳ x y z = to (distribʳ x y z) distrib : ∀ {∙ ◦} → _DistributesOver_ ≈₁ ∙ ◦ → _DistributesOver_ ≈₂ ∙ ◦ -distrib {∙} {◦} = Prod.map (distribˡ {∙} {◦}) (distribʳ {∙} {◦}) +distrib {∙} {◦} = Product.map (distribˡ {∙} {◦}) (distribʳ {∙} {◦}) ------------------------------------------------------------------------ -- Structures @@ -92,7 +92,7 @@ isSelectiveMagma S = record isMonoid : ∀ {∙ ε} → IsMonoid ≈₁ ∙ ε → IsMonoid ≈₂ ∙ ε isMonoid S = record { isSemigroup = isSemigroup S.isSemigroup - ; identity = Prod.map (to ∘_) (to ∘_) S.identity + ; identity = Product.map (to ∘_) (to ∘_) S.identity } where module S = IsMonoid S isCommutativeMonoid : ∀ {∙ ε} → @@ -113,7 +113,7 @@ isIdempotentCommutativeMonoid {∙} S = record isGroup : ∀ {∙ ε ⁻¹} → IsGroup ≈₁ ∙ ε ⁻¹ → IsGroup ≈₂ ∙ ε ⁻¹ isGroup S = record { isMonoid = isMonoid S.isMonoid - ; inverse = Prod.map (to ∘_) (to ∘_) S.inverse + ; inverse = Product.map (to ∘_) (to ∘_) S.inverse ; ⁻¹-cong = cong₁ S.⁻¹-cong } where module S = IsGroup S @@ -141,7 +141,7 @@ isSemiringWithoutOne {+} {*} S = record ; *-cong = cong₂ S.*-cong ; *-assoc = assoc {*} S.*-assoc ; distrib = distrib {*} {+} S.distrib - ; zero = Prod.map (to ∘_) (to ∘_) S.zero + ; zero = Product.map (to ∘_) (to ∘_) S.zero } where module S = IsSemiringWithoutOne S isCommutativeSemiringWithoutOne : ∀ {+ * 0#} → diff --git a/src/Algebra/Lattice/Construct/Subst/Equality.agda b/src/Algebra/Lattice/Construct/Subst/Equality.agda index e756319aaf..8a498afbb2 100644 --- a/src/Algebra/Lattice/Construct/Subst/Equality.agda +++ b/src/Algebra/Lattice/Construct/Subst/Equality.agda @@ -12,7 +12,7 @@ open import Algebra.Core using (Op₂) open import Algebra.Definitions open import Algebra.Lattice.Structures -open import Data.Product.Base as Prod +open import Data.Product.Base using (_,_) open import Function.Base open import Relation.Binary.Core diff --git a/src/Algebra/Lattice/Morphism/LatticeMonomorphism.agda b/src/Algebra/Lattice/Morphism/LatticeMonomorphism.agda index 9eb1d4da24..9e16a46c4a 100644 --- a/src/Algebra/Lattice/Morphism/LatticeMonomorphism.agda +++ b/src/Algebra/Lattice/Morphism/LatticeMonomorphism.agda @@ -18,7 +18,7 @@ import Algebra.Lattice.Properties.Lattice as LatticeProperties open import Data.Product.Base using (_,_; map) open import Relation.Binary.Bundles using (Setoid) import Relation.Binary.Morphism.RelMonomorphism as RelMonomorphisms -import Relation.Binary.Reasoning.Setoid as SetoidReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning module Algebra.Lattice.Morphism.LatticeMonomorphism {a b ℓ₁ ℓ₂} {L₁ : RawLattice a ℓ₁} {L₂ : RawLattice b ℓ₂} {⟦_⟧} @@ -73,7 +73,7 @@ module _ (⊔-⊓-isLattice : IsLattice _≈₂_ _⊔_ _⊓_) where setoid : Setoid b ℓ₂ setoid = record { isEquivalence = isEquivalence } - open SetoidReasoning setoid + open ≈-Reasoning setoid ∨-absorbs-∧ : _Absorbs_ _≈₁_ _∨_ _∧_ ∨-absorbs-∧ x y = injective (begin diff --git a/src/Algebra/Module/Bundles.agda b/src/Algebra/Module/Bundles.agda index f36e5bccd0..33e6bfa731 100644 --- a/src/Algebra/Module/Bundles.agda +++ b/src/Algebra/Module/Bundles.agda @@ -28,6 +28,7 @@ module Algebra.Module.Bundles where open import Algebra.Bundles open import Algebra.Core open import Algebra.Definitions using (Involutive) +import Algebra.Module.Bundles.Raw as Raw open import Algebra.Module.Core open import Algebra.Module.Structures open import Algebra.Module.Definitions @@ -36,12 +37,22 @@ open import Function.Base open import Level open import Relation.Binary.Core using (Rel) open import Relation.Nullary using (¬_) -import Relation.Binary.Reasoning.Setoid as SetR +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning private variable r ℓr s ℓs : Level +------------------------------------------------------------------------ +-- Re-export definitions of 'raw' bundles + +open Raw public + using ( RawLeftSemimodule; RawLeftModule + ; RawRightSemimodule; RawRightModule + ; RawBisemimodule; RawBimodule + ; RawSemimodule; RawModule + ) + ------------------------------------------------------------------------ -- Left modules ------------------------------------------------------------------------ @@ -62,10 +73,6 @@ record LeftSemimodule (semiring : Semiring r ℓr) m ℓm 0ᴹ : Carrierᴹ isLeftSemimodule : IsLeftSemimodule semiring _≈ᴹ_ _+ᴹ_ 0ᴹ _*ₗ_ - infix 4 _≉ᴹ_ - _≉ᴹ_ : Rel Carrierᴹ _ - a ≉ᴹ b = ¬ (a ≈ᴹ b) - open IsLeftSemimodule isLeftSemimodule public +ᴹ-commutativeMonoid : CommutativeMonoid m ℓm @@ -80,8 +87,17 @@ record LeftSemimodule (semiring : Semiring r ℓr) m ℓm ; magma to +ᴹ-magma ; rawMagma to +ᴹ-rawMagma ; rawMonoid to +ᴹ-rawMonoid + ; _≉_ to _≉ᴹ_ ) + rawLeftSemimodule : RawLeftSemimodule Carrier m ℓm + rawLeftSemimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ₗ_ = _*ₗ_ + ; 0ᴹ = 0ᴹ + } + record LeftModule (ring : Ring r ℓr) m ℓm : Set (r ⊔ ℓr ⊔ suc (m ⊔ ℓm)) where open Ring ring @@ -106,7 +122,7 @@ record LeftModule (ring : Ring r ℓr) m ℓm : Set (r ⊔ ℓr ⊔ suc (m ⊔ open LeftSemimodule leftSemimodule public using ( +ᴹ-commutativeMonoid; +ᴹ-monoid; +ᴹ-semigroup; +ᴹ-magma - ; +ᴹ-rawMagma; +ᴹ-rawMonoid; _≉ᴹ_) + ; +ᴹ-rawMagma; +ᴹ-rawMonoid; rawLeftSemimodule; _≉ᴹ_) +ᴹ-abelianGroup : AbelianGroup m ℓm +ᴹ-abelianGroup = record { isAbelianGroup = +ᴹ-isAbelianGroup } @@ -114,6 +130,15 @@ record LeftModule (ring : Ring r ℓr) m ℓm : Set (r ⊔ ℓr ⊔ suc (m ⊔ open AbelianGroup +ᴹ-abelianGroup public using () renaming (group to +ᴹ-group; rawGroup to +ᴹ-rawGroup) + rawLeftModule : RawLeftModule Carrier m ℓm + rawLeftModule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ₗ_ = _*ₗ_ + ; 0ᴹ = 0ᴹ + ; -ᴹ_ = -ᴹ_ + } + ------------------------------------------------------------------------ -- Right modules ------------------------------------------------------------------------ @@ -134,10 +159,6 @@ record RightSemimodule (semiring : Semiring r ℓr) m ℓm 0ᴹ : Carrierᴹ isRightSemimodule : IsRightSemimodule semiring _≈ᴹ_ _+ᴹ_ 0ᴹ _*ᵣ_ - infix 4 _≉ᴹ_ - _≉ᴹ_ : Rel Carrierᴹ _ - a ≉ᴹ b = ¬ (a ≈ᴹ b) - open IsRightSemimodule isRightSemimodule public +ᴹ-commutativeMonoid : CommutativeMonoid m ℓm @@ -152,8 +173,17 @@ record RightSemimodule (semiring : Semiring r ℓr) m ℓm ; magma to +ᴹ-magma ; rawMagma to +ᴹ-rawMagma ; rawMonoid to +ᴹ-rawMonoid + ; _≉_ to _≉ᴹ_ ) + rawRightSemimodule : RawRightSemimodule Carrier m ℓm + rawRightSemimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ᵣ_ = _*ᵣ_ + ; 0ᴹ = 0ᴹ + } + record RightModule (ring : Ring r ℓr) m ℓm : Set (r ⊔ ℓr ⊔ suc (m ⊔ ℓm)) where open Ring ring @@ -178,7 +208,7 @@ record RightModule (ring : Ring r ℓr) m ℓm : Set (r ⊔ ℓr ⊔ suc (m ⊔ open RightSemimodule rightSemimodule public using ( +ᴹ-commutativeMonoid; +ᴹ-monoid; +ᴹ-semigroup; +ᴹ-magma - ; +ᴹ-rawMagma; +ᴹ-rawMonoid; _≉ᴹ_) + ; +ᴹ-rawMagma; +ᴹ-rawMonoid; rawRightSemimodule; _≉ᴹ_) +ᴹ-abelianGroup : AbelianGroup m ℓm +ᴹ-abelianGroup = record { isAbelianGroup = +ᴹ-isAbelianGroup } @@ -186,6 +216,15 @@ record RightModule (ring : Ring r ℓr) m ℓm : Set (r ⊔ ℓr ⊔ suc (m ⊔ open AbelianGroup +ᴹ-abelianGroup public using () renaming (group to +ᴹ-group; rawGroup to +ᴹ-rawGroup) + rawRightModule : RawRightModule Carrier m ℓm + rawRightModule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ᵣ_ = _*ᵣ_ + ; 0ᴹ = 0ᴹ + ; -ᴹ_ = -ᴹ_ + } + ------------------------------------------------------------------------ -- Bimodules ------------------------------------------------------------------------ @@ -220,7 +259,19 @@ record Bisemimodule (R-semiring : Semiring r ℓr) (S-semiring : Semiring s ℓs open LeftSemimodule leftSemimodule public using ( +ᴹ-commutativeMonoid; +ᴹ-monoid; +ᴹ-semigroup; +ᴹ-magma; +ᴹ-rawMagma - ; +ᴹ-rawMonoid; _≉ᴹ_) + ; +ᴹ-rawMonoid; rawLeftSemimodule; _≉ᴹ_) + + open RightSemimodule rightSemimodule public + using ( rawRightSemimodule ) + + rawBisemimodule : RawBisemimodule R.Carrier S.Carrier m ℓm + rawBisemimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ₗ_ = _*ₗ_ + ; _*ᵣ_ = _*ᵣ_ + ; 0ᴹ = 0ᴹ + } record Bimodule (R-ring : Ring r ℓr) (S-ring : Ring s ℓs) m ℓm : Set (r ⊔ s ⊔ ℓr ⊔ ℓs ⊔ suc (m ⊔ ℓm)) where @@ -254,13 +305,26 @@ record Bimodule (R-ring : Ring r ℓr) (S-ring : Ring s ℓs) m ℓm open LeftModule leftModule public using ( +ᴹ-abelianGroup; +ᴹ-commutativeMonoid; +ᴹ-group; +ᴹ-monoid ; +ᴹ-semigroup; +ᴹ-magma; +ᴹ-rawMagma; +ᴹ-rawMonoid; +ᴹ-rawGroup - ; _≉ᴹ_) + ; rawLeftSemimodule; rawLeftModule; _≉ᴹ_) + + open RightModule rightModule public + using ( rawRightSemimodule; rawRightModule ) bisemimodule : Bisemimodule R.semiring S.semiring m ℓm bisemimodule = record { isBisemimodule = isBisemimodule } open Bisemimodule bisemimodule public - using (leftSemimodule; rightSemimodule) + using (leftSemimodule; rightSemimodule; rawBisemimodule) + + rawBimodule : RawBimodule R.Carrier S.Carrier m ℓm + rawBimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ₗ_ = _*ₗ_ + ; _*ᵣ_ = _*ᵣ_ + ; 0ᴹ = 0ᴹ + ; -ᴹ_ = -ᴹ_ + } ------------------------------------------------------------------------ -- Modules over commutative structures @@ -296,9 +360,11 @@ record Semimodule (commutativeSemiring : CommutativeSemiring r ℓr) m ℓm open Bisemimodule bisemimodule public using ( leftSemimodule; rightSemimodule ; +ᴹ-commutativeMonoid; +ᴹ-monoid; +ᴹ-semigroup; +ᴹ-magma - ; +ᴹ-rawMagma; +ᴹ-rawMonoid; _≉ᴹ_) + ; +ᴹ-rawMagma; +ᴹ-rawMonoid; rawLeftSemimodule; rawRightSemimodule + ; rawBisemimodule; _≉ᴹ_ + ) - open SetR ≈ᴹ-setoid + open ≈-Reasoning ≈ᴹ-setoid *ₗ-comm : L.Commutative _*ₗ_ *ₗ-comm x y m = begin @@ -314,6 +380,9 @@ record Semimodule (commutativeSemiring : CommutativeSemiring r ℓr) m ℓm m *ᵣ (y * x) ≈⟨ ≈ᴹ-sym (*ᵣ-assoc m y x) ⟩ m *ᵣ y *ᵣ x ∎ + rawSemimodule : RawSemimodule Carrier m ℓm + rawSemimodule = rawBisemimodule + record Module (commutativeRing : CommutativeRing r ℓr) m ℓm : Set (r ⊔ ℓr ⊔ suc (m ⊔ ℓm)) where open CommutativeRing commutativeRing @@ -343,9 +412,13 @@ record Module (commutativeRing : CommutativeRing r ℓr) m ℓm using ( leftModule; rightModule; leftSemimodule; rightSemimodule ; +ᴹ-abelianGroup; +ᴹ-group; +ᴹ-commutativeMonoid; +ᴹ-monoid ; +ᴹ-semigroup; +ᴹ-magma ; +ᴹ-rawMonoid; +ᴹ-rawMagma - ; +ᴹ-rawGroup; _≉ᴹ_) + ; +ᴹ-rawGroup; rawLeftSemimodule; rawLeftModule; rawRightSemimodule + ; rawRightModule; rawBisemimodule; rawBimodule; _≉ᴹ_) semimodule : Semimodule commutativeSemiring m ℓm semimodule = record { isSemimodule = isSemimodule } - open Semimodule semimodule public using (*ₗ-comm; *ᵣ-comm) + open Semimodule semimodule public using (*ₗ-comm; *ᵣ-comm; rawSemimodule) + + rawModule : RawModule Carrier m ℓm + rawModule = rawBimodule diff --git a/src/Algebra/Module/Bundles/Raw.agda b/src/Algebra/Module/Bundles/Raw.agda new file mode 100644 index 0000000000..b9cc108915 --- /dev/null +++ b/src/Algebra/Module/Bundles/Raw.agda @@ -0,0 +1,251 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Definitions of 'raw' bundles for module-like algebraic structures +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Algebra.Module.Bundles.Raw where + +open import Algebra.Bundles.Raw +open import Algebra.Core +open import Algebra.Module.Core +open import Level +open import Relation.Nullary.Negation.Core using (¬_) +open import Relation.Binary.Core using (Rel) + +private + variable + r ℓr s ℓs : Level + +------------------------------------------------------------------------ +-- Raw left modules +------------------------------------------------------------------------ + +record RawLeftSemimodule (R : Set r) m ℓm : Set (r ⊔ suc (m ⊔ ℓm)) where + infixr 7 _*ₗ_ + infixl 6 _+ᴹ_ + infix 4 _≈ᴹ_ + + field + Carrierᴹ : Set m + _≈ᴹ_ : Rel Carrierᴹ ℓm + _+ᴹ_ : Op₂ Carrierᴹ + _*ₗ_ : Opₗ R Carrierᴹ + 0ᴹ : Carrierᴹ + + +ᴹ-rawMonoid : RawMonoid m ℓm + +ᴹ-rawMonoid = record + { _≈_ = _≈ᴹ_ + ; _∙_ = _+ᴹ_ + ; ε = 0ᴹ + } + + open RawMonoid +ᴹ-rawMonoid public + using () + renaming (rawMagma to +ᴹ-rawMagma; _≉_ to _≉ᴹ_) + +record RawLeftModule (R : Set r) m ℓm : Set (r ⊔ suc (m ⊔ ℓm)) where + infix 8 -ᴹ_ + infixr 7 _*ₗ_ + infixl 6 _+ᴹ_ + infix 4 _≈ᴹ_ + + field + Carrierᴹ : Set m + _≈ᴹ_ : Rel Carrierᴹ ℓm + _+ᴹ_ : Op₂ Carrierᴹ + _*ₗ_ : Opₗ R Carrierᴹ + 0ᴹ : Carrierᴹ + -ᴹ_ : Op₁ Carrierᴹ + + rawLeftSemimodule : RawLeftSemimodule R m ℓm + rawLeftSemimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ₗ_ = _*ₗ_ + ; 0ᴹ = 0ᴹ + } + + open RawLeftSemimodule rawLeftSemimodule public + using (+ᴹ-rawMagma; +ᴹ-rawMonoid; _≉ᴹ_) + + +ᴹ-rawGroup : RawGroup m ℓm + +ᴹ-rawGroup = record + { _≈_ = _≈ᴹ_ + ; _∙_ = _+ᴹ_ + ; ε = 0ᴹ + ; _⁻¹ = -ᴹ_ + } + +------------------------------------------------------------------------ +-- Raw right modules +------------------------------------------------------------------------ + +record RawRightSemimodule (R : Set r) m ℓm : Set (r ⊔ suc (m ⊔ ℓm)) where + infixl 7 _*ᵣ_ + infixl 6 _+ᴹ_ + infix 4 _≈ᴹ_ + + field + Carrierᴹ : Set m + _≈ᴹ_ : Rel Carrierᴹ ℓm + _+ᴹ_ : Op₂ Carrierᴹ + _*ᵣ_ : Opᵣ R Carrierᴹ + 0ᴹ : Carrierᴹ + + +ᴹ-rawMonoid : RawMonoid m ℓm + +ᴹ-rawMonoid = record + { _≈_ = _≈ᴹ_ + ; _∙_ = _+ᴹ_ + ; ε = 0ᴹ + } + + open RawMonoid +ᴹ-rawMonoid public + using () + renaming (rawMagma to +ᴹ-rawMagma; _≉_ to _≉ᴹ_) + +record RawRightModule (R : Set r) m ℓm : Set (r ⊔ suc (m ⊔ ℓm)) where + infix 8 -ᴹ_ + infixl 7 _*ᵣ_ + infixl 6 _+ᴹ_ + infix 4 _≈ᴹ_ + + field + Carrierᴹ : Set m + _≈ᴹ_ : Rel Carrierᴹ ℓm + _+ᴹ_ : Op₂ Carrierᴹ + _*ᵣ_ : Opᵣ R Carrierᴹ + 0ᴹ : Carrierᴹ + -ᴹ_ : Op₁ Carrierᴹ + + rawRightSemimodule : RawRightSemimodule R m ℓm + rawRightSemimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ᵣ_ = _*ᵣ_ + ; 0ᴹ = 0ᴹ + } + + open RawRightSemimodule rawRightSemimodule public + using (+ᴹ-rawMagma; +ᴹ-rawMonoid; _≉ᴹ_) + + +ᴹ-rawGroup : RawGroup m ℓm + +ᴹ-rawGroup = record + { _≈_ = _≈ᴹ_ + ; _∙_ = _+ᴹ_ + ; ε = 0ᴹ + ; _⁻¹ = -ᴹ_ + } + +------------------------------------------------------------------------ +-- Bimodules +------------------------------------------------------------------------ + +record RawBisemimodule (R : Set r) (S : Set s) m ℓm : Set (r ⊔ s ⊔ suc (m ⊔ ℓm)) where + infixr 7 _*ₗ_ + infixl 7 _*ᵣ_ + infixl 6 _+ᴹ_ + infix 4 _≈ᴹ_ + + field + Carrierᴹ : Set m + _≈ᴹ_ : Rel Carrierᴹ ℓm + _+ᴹ_ : Op₂ Carrierᴹ + _*ₗ_ : Opₗ R Carrierᴹ + _*ᵣ_ : Opᵣ S Carrierᴹ + 0ᴹ : Carrierᴹ + + rawLeftSemimodule : RawLeftSemimodule R m ℓm + rawLeftSemimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ₗ_ = _*ₗ_ + ; 0ᴹ = 0ᴹ + } + + rawRightSemimodule : RawRightSemimodule S m ℓm + rawRightSemimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ᵣ_ = _*ᵣ_ + ; 0ᴹ = 0ᴹ + } + + open RawLeftSemimodule rawLeftSemimodule public + using (+ᴹ-rawMagma; +ᴹ-rawMonoid; _≉ᴹ_) + +record RawBimodule (R : Set r) (S : Set s) m ℓm : Set (r ⊔ s ⊔ suc (m ⊔ ℓm)) where + infix 8 -ᴹ_ + infixr 7 _*ₗ_ + infixl 7 _*ᵣ_ + infixl 6 _+ᴹ_ + infix 4 _≈ᴹ_ + + field + Carrierᴹ : Set m + _≈ᴹ_ : Rel Carrierᴹ ℓm + _+ᴹ_ : Op₂ Carrierᴹ + _*ₗ_ : Opₗ R Carrierᴹ + _*ᵣ_ : Opᵣ S Carrierᴹ + 0ᴹ : Carrierᴹ + -ᴹ_ : Op₁ Carrierᴹ + + rawLeftModule : RawLeftModule R m ℓm + rawLeftModule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ₗ_ = _*ₗ_ + ; 0ᴹ = 0ᴹ + ; -ᴹ_ = -ᴹ_ + } + + rawRightModule : RawRightModule S m ℓm + rawRightModule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ᵣ_ = _*ᵣ_ + ; 0ᴹ = 0ᴹ + ; -ᴹ_ = -ᴹ_ + } + + rawBisemimodule : RawBisemimodule R S m ℓm + rawBisemimodule = record + { _≈ᴹ_ = _≈ᴹ_ + ; _+ᴹ_ = _+ᴹ_ + ; _*ₗ_ = _*ₗ_ + ; _*ᵣ_ = _*ᵣ_ + ; 0ᴹ = 0ᴹ + } + + open RawBisemimodule rawBisemimodule public + using (+ᴹ-rawMagma; +ᴹ-rawMonoid; rawLeftSemimodule; rawRightSemimodule; _≉ᴹ_) + + open RawLeftModule rawLeftModule public + using (+ᴹ-rawGroup) + +------------------------------------------------------------------------ +-- Modules over commutative structures +------------------------------------------------------------------------ + +RawSemimodule : ∀ (R : Set r) m ℓm → Set (r ⊔ suc (m ⊔ ℓm)) +RawSemimodule R = RawBisemimodule R R + +module RawSemimodule {R : Set r} {m ℓm} (M : RawSemimodule R m ℓm) where + open RawBisemimodule M public + + rawBisemimodule : RawBisemimodule R R m ℓm + rawBisemimodule = M + +RawModule : ∀ (R : Set r) m ℓm → Set (r ⊔ suc(m ⊔ ℓm)) +RawModule R = RawBimodule R R + +module RawModule {R : Set r} {m ℓm} (M : RawModule R m ℓm) where + open RawBimodule M public + + rawBimodule : RawBimodule R R m ℓm + rawBimodule = M + + rawSemimodule : RawSemimodule R m ℓm + rawSemimodule = rawBisemimodule diff --git a/src/Algebra/Module/Consequences.agda b/src/Algebra/Module/Consequences.agda index 9acd6d6413..2ee2c98d7e 100644 --- a/src/Algebra/Module/Consequences.agda +++ b/src/Algebra/Module/Consequences.agda @@ -16,7 +16,7 @@ open import Function.Base using (flip) open import Level using (Level) open import Relation.Binary.Core using (Rel) open import Relation.Binary.Bundles using (Setoid) -import Relation.Binary.Reasoning.Setoid as Rea +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning private variable @@ -27,7 +27,7 @@ private module _ (_≈ᴬ_ : Rel {a} A ℓa) (S : Setoid c ℓ) where open Setoid S - open Rea S + open ≈-Reasoning S open Defs _≈ᴬ_ private diff --git a/src/Algebra/Module/Construct/DirectProduct.agda b/src/Algebra/Module/Construct/DirectProduct.agda index 429a2d11ca..473e240801 100644 --- a/src/Algebra/Module/Construct/DirectProduct.agda +++ b/src/Algebra/Module/Construct/DirectProduct.agda @@ -14,7 +14,7 @@ module Algebra.Module.Construct.DirectProduct where open import Algebra.Bundles open import Algebra.Construct.DirectProduct open import Algebra.Module.Bundles -open import Data.Product.Base using (map; _,_; proj₁; proj₂) +open import Data.Product.Base using (map; zip; _,_; proj₁; proj₂) open import Data.Product.Relation.Binary.Pointwise.NonDependent open import Level @@ -22,6 +22,83 @@ private variable r s ℓr ℓs m m′ ℓm ℓm′ : Level +------------------------------------------------------------------------ +-- Raw bundles + +rawLeftSemimodule : {R : Set r} → + RawLeftSemimodule R m ℓm → + RawLeftSemimodule R m′ ℓm′ → + RawLeftSemimodule R (m ⊔ m′) (ℓm ⊔ ℓm′) +rawLeftSemimodule M N = record + { _≈ᴹ_ = Pointwise M._≈ᴹ_ N._≈ᴹ_ + ; _+ᴹ_ = zip M._+ᴹ_ N._+ᴹ_ + ; _*ₗ_ = λ r → map (r M.*ₗ_) (r N.*ₗ_) + ; 0ᴹ = M.0ᴹ , N.0ᴹ + } where module M = RawLeftSemimodule M; module N = RawLeftSemimodule N + +rawLeftModule : {R : Set r} → + RawLeftModule R m ℓm → + RawLeftModule R m′ ℓm′ → + RawLeftModule R (m ⊔ m′) (ℓm ⊔ ℓm′) +rawLeftModule M N = record + { RawLeftSemimodule (rawLeftSemimodule M.rawLeftSemimodule N.rawLeftSemimodule) + ; -ᴹ_ = map M.-ᴹ_ N.-ᴹ_ + } where module M = RawLeftModule M; module N = RawLeftModule N + + +rawRightSemimodule : {R : Set r} → + RawRightSemimodule R m ℓm → + RawRightSemimodule R m′ ℓm′ → + RawRightSemimodule R (m ⊔ m′) (ℓm ⊔ ℓm′) +rawRightSemimodule M N = record + { _≈ᴹ_ = Pointwise M._≈ᴹ_ N._≈ᴹ_ + ; _+ᴹ_ = zip M._+ᴹ_ N._+ᴹ_ + ; _*ᵣ_ = λ mn r → map (M._*ᵣ r) (N._*ᵣ r) mn + ; 0ᴹ = M.0ᴹ , N.0ᴹ + } where module M = RawRightSemimodule M; module N = RawRightSemimodule N + +rawRightModule : {R : Set r} → + RawRightModule R m ℓm → + RawRightModule R m′ ℓm′ → + RawRightModule R (m ⊔ m′) (ℓm ⊔ ℓm′) +rawRightModule M N = record + { RawRightSemimodule (rawRightSemimodule M.rawRightSemimodule N.rawRightSemimodule) + ; -ᴹ_ = map M.-ᴹ_ N.-ᴹ_ + } where module M = RawRightModule M; module N = RawRightModule N + +rawBisemimodule : {R : Set r} {S : Set s} → + RawBisemimodule R S m ℓm → + RawBisemimodule R S m′ ℓm′ → + RawBisemimodule R S (m ⊔ m′) (ℓm ⊔ ℓm′) +rawBisemimodule M N = record + { _≈ᴹ_ = Pointwise M._≈ᴹ_ N._≈ᴹ_ + ; _+ᴹ_ = zip M._+ᴹ_ N._+ᴹ_ + ; _*ₗ_ = λ r → map (r M.*ₗ_) (r N.*ₗ_) + ; _*ᵣ_ = λ mn r → map (M._*ᵣ r) (N._*ᵣ r) mn + ; 0ᴹ = M.0ᴹ , N.0ᴹ + } where module M = RawBisemimodule M; module N = RawBisemimodule N + +rawBimodule : {R : Set r} {S : Set s} → + RawBimodule R S m ℓm → + RawBimodule R S m′ ℓm′ → + RawBimodule R S (m ⊔ m′) (ℓm ⊔ ℓm′) +rawBimodule M N = record + { RawBisemimodule (rawBisemimodule M.rawBisemimodule N.rawBisemimodule) + ; -ᴹ_ = map M.-ᴹ_ N.-ᴹ_ + } where module M = RawBimodule M; module N = RawBimodule N + +rawSemimodule : {R : Set r} → + RawSemimodule R m ℓm → + RawSemimodule R m′ ℓm′ → + RawSemimodule R (m ⊔ m′) (ℓm ⊔ ℓm′) +rawSemimodule M N = rawBisemimodule M N + +rawModule : {R : Set r} → + RawModule R m ℓm → + RawModule R m′ ℓm′ → + RawModule R (m ⊔ m′) (ℓm ⊔ ℓm′) +rawModule M N = rawBimodule M N + ------------------------------------------------------------------------ -- Bundles diff --git a/src/Algebra/Module/Construct/TensorUnit.agda b/src/Algebra/Module/Construct/TensorUnit.agda index c76f2a39ba..6c97b67de0 100644 --- a/src/Algebra/Module/Construct/TensorUnit.agda +++ b/src/Algebra/Module/Construct/TensorUnit.agda @@ -19,6 +19,62 @@ private variable c ℓ : Level +------------------------------------------------------------------------ +-- Raw bundles + +rawLeftSemimodule : {R : RawSemiring c ℓ} → RawLeftSemimodule _ c ℓ +rawLeftSemimodule {R = R} = record + { _≈ᴹ_ = _≈_ + ; _+ᴹ_ = _+_ + ; _*ₗ_ = _*_ + ; 0ᴹ = 0# + } where open RawSemiring R + +rawLeftModule : {R : RawRing c ℓ} → RawLeftModule _ c ℓ +rawLeftModule {R = R} = record + { RawLeftSemimodule (rawLeftSemimodule {R = rawSemiring}) + ; -ᴹ_ = -_ + } where open RawRing R + +rawRightSemimodule : {R : RawSemiring c ℓ} → RawRightSemimodule _ c ℓ +rawRightSemimodule {R = R} = record + { _≈ᴹ_ = _≈_ + ; _+ᴹ_ = _+_ + ; _*ᵣ_ = _*_ + ; 0ᴹ = 0# + } where open RawSemiring R + +rawRightModule : {R : RawRing c ℓ} → RawRightModule _ c ℓ +rawRightModule {R = R} = record + { RawRightSemimodule (rawRightSemimodule {R = rawSemiring}) + ; -ᴹ_ = -_ + } where open RawRing R + +rawBisemimodule : {R : RawSemiring c ℓ} → RawBisemimodule _ _ c ℓ +rawBisemimodule {R = R} = record + { _≈ᴹ_ = _≈_ + ; _+ᴹ_ = _+_ + ; _*ₗ_ = _*_ + ; _*ᵣ_ = _*_ + ; 0ᴹ = 0# + } where open RawSemiring R + +rawBimodule : {R : RawRing c ℓ} → RawBimodule _ _ c ℓ +rawBimodule {R = R} = record + { RawBisemimodule (rawBisemimodule {R = rawSemiring}) + ; -ᴹ_ = -_ + } where open RawRing R + + +rawSemimodule : {R : RawSemiring c ℓ} → RawSemimodule _ c ℓ +rawSemimodule {R = R} = rawBisemimodule {R = R} + +rawModule : {R : RawRing c ℓ} → RawModule _ c ℓ +rawModule {R = R} = rawBimodule {R = R} + +------------------------------------------------------------------------ +-- Bundles + leftSemimodule : {R : Semiring c ℓ} → LeftSemimodule R c ℓ leftSemimodule {R = semiring} = record { Carrierᴹ = Carrier diff --git a/src/Algebra/Module/Construct/Zero.agda b/src/Algebra/Module/Construct/Zero.agda index 624cd04598..256f70e8c0 100644 --- a/src/Algebra/Module/Construct/Zero.agda +++ b/src/Algebra/Module/Construct/Zero.agda @@ -35,14 +35,33 @@ module ℤero where _≈ᴹ_ : Rel Carrierᴹ ℓ _ ≈ᴹ _ = ⊤ -open ℤero - ------------------------------------------------------------------------ --- Raw bundles NOT YET IMPLEMENTED issue #1828 -{- -rawModule : RawModule c ℓ +-- Raw bundles + +rawLeftSemimodule : {R : Set r} → RawLeftSemimodule R c ℓ +rawLeftSemimodule = record { ℤero } + +rawLeftModule : {R : Set r} → RawLeftModule R c ℓ +rawLeftModule = record { ℤero } + +rawRightSemimodule : {R : Set r} → RawRightSemimodule R c ℓ +rawRightSemimodule = record { ℤero } + +rawRightModule : {R : Set r} → RawRightModule R c ℓ +rawRightModule = record { ℤero } + +rawBisemimodule : {R : Set r} {S : Set s} → RawBisemimodule R S c ℓ +rawBisemimodule = record { ℤero } + +rawBimodule : {R : Set r} {S : Set s} → RawBimodule R S c ℓ +rawBimodule = record { ℤero } + +rawSemimodule : {R : Set r} → RawSemimodule R c ℓ +rawSemimodule = record { ℤero } + +rawModule : {R : Set r} → RawModule R c ℓ rawModule = record { ℤero } --} + ------------------------------------------------------------------------ -- Bundles diff --git a/src/Algebra/Module/Morphism/Construct/Composition.agda b/src/Algebra/Module/Morphism/Construct/Composition.agda index 8c39496dd4..ddeeef6893 100644 --- a/src/Algebra/Module/Morphism/Construct/Composition.agda +++ b/src/Algebra/Module/Morphism/Construct/Composition.agda @@ -8,24 +8,25 @@ module Algebra.Module.Morphism.Construct.Composition where -open import Algebra.Bundles -open import Algebra.Module.Bundles +open import Algebra.Module.Bundles.Raw open import Algebra.Module.Morphism.Structures open import Algebra.Morphism.Construct.Composition open import Function.Base using (_∘_) import Function.Construct.Composition as Func open import Level using (Level) +open import Relation.Binary.Definitions using (Transitive) private variable - r ℓr s ℓs m₁ ℓm₁ m₂ ℓm₂ m₃ ℓm₃ : Level + r s m₁ ℓm₁ m₂ ℓm₂ m₃ ℓm₃ : Level module _ - {semiring : Semiring r ℓr} - {M₁ : LeftSemimodule semiring m₁ ℓm₁} - {M₂ : LeftSemimodule semiring m₂ ℓm₂} - {M₃ : LeftSemimodule semiring m₃ ℓm₃} - (open LeftSemimodule) + {R : Set r} + {M₁ : RawLeftSemimodule R m₁ ℓm₁} + {M₂ : RawLeftSemimodule R m₂ ℓm₂} + {M₃ : RawLeftSemimodule R m₃ ℓm₃} + (open RawLeftSemimodule) + (≈ᴹ₃-trans : Transitive (_≈ᴹ_ M₃)) {f : Carrierᴹ M₁ → Carrierᴹ M₂} {g : Carrierᴹ M₂ → Carrierᴹ M₃} where @@ -34,8 +35,8 @@ module _ IsLeftSemimoduleHomomorphism M₂ M₃ g → IsLeftSemimoduleHomomorphism M₁ M₃ (g ∘ f) isLeftSemimoduleHomomorphism f-homo g-homo = record - { +ᴹ-isMonoidHomomorphism = isMonoidHomomorphism (≈ᴹ-trans M₃) F.+ᴹ-isMonoidHomomorphism G.+ᴹ-isMonoidHomomorphism - ; *ₗ-homo = λ r x → ≈ᴹ-trans M₃ (G.⟦⟧-cong (F.*ₗ-homo r x)) (G.*ₗ-homo r (f x)) + { +ᴹ-isMonoidHomomorphism = isMonoidHomomorphism ≈ᴹ₃-trans F.+ᴹ-isMonoidHomomorphism G.+ᴹ-isMonoidHomomorphism + ; *ₗ-homo = λ r x → ≈ᴹ₃-trans (G.⟦⟧-cong (F.*ₗ-homo r x)) (G.*ₗ-homo r (f x)) } where module F = IsLeftSemimoduleHomomorphism f-homo; module G = IsLeftSemimoduleHomomorphism g-homo isLeftSemimoduleMonomorphism : IsLeftSemimoduleMonomorphism M₁ M₂ f → @@ -55,11 +56,12 @@ module _ } where module F = IsLeftSemimoduleIsomorphism f-iso; module G = IsLeftSemimoduleIsomorphism g-iso module _ - {ring : Ring r ℓr} - {M₁ : LeftModule ring m₁ ℓm₁} - {M₂ : LeftModule ring m₂ ℓm₂} - {M₃ : LeftModule ring m₃ ℓm₃} - (open LeftModule) + {R : Set r} + {M₁ : RawLeftModule R m₁ ℓm₁} + {M₂ : RawLeftModule R m₂ ℓm₂} + {M₃ : RawLeftModule R m₃ ℓm₃} + (open RawLeftModule) + (≈ᴹ₃-trans : Transitive (_≈ᴹ_ M₃)) {f : Carrierᴹ M₁ → Carrierᴹ M₂} {g : Carrierᴹ M₂ → Carrierᴹ M₃} where @@ -68,8 +70,8 @@ module _ IsLeftModuleHomomorphism M₂ M₃ g → IsLeftModuleHomomorphism M₁ M₃ (g ∘ f) isLeftModuleHomomorphism f-homo g-homo = record - { +ᴹ-isGroupHomomorphism = isGroupHomomorphism (≈ᴹ-trans M₃) F.+ᴹ-isGroupHomomorphism G.+ᴹ-isGroupHomomorphism - ; *ₗ-homo = λ r x → ≈ᴹ-trans M₃ (G.⟦⟧-cong (F.*ₗ-homo r x)) (G.*ₗ-homo r (f x)) + { +ᴹ-isGroupHomomorphism = isGroupHomomorphism ≈ᴹ₃-trans F.+ᴹ-isGroupHomomorphism G.+ᴹ-isGroupHomomorphism + ; *ₗ-homo = λ r x → ≈ᴹ₃-trans (G.⟦⟧-cong (F.*ₗ-homo r x)) (G.*ₗ-homo r (f x)) } where module F = IsLeftModuleHomomorphism f-homo; module G = IsLeftModuleHomomorphism g-homo isLeftModuleMonomorphism : IsLeftModuleMonomorphism M₁ M₂ f → @@ -89,11 +91,12 @@ module _ } where module F = IsLeftModuleIsomorphism f-iso; module G = IsLeftModuleIsomorphism g-iso module _ - {semiring : Semiring r ℓr} - {M₁ : RightSemimodule semiring m₁ ℓm₁} - {M₂ : RightSemimodule semiring m₂ ℓm₂} - {M₃ : RightSemimodule semiring m₃ ℓm₃} - (open RightSemimodule) + {R : Set r} + {M₁ : RawRightSemimodule R m₁ ℓm₁} + {M₂ : RawRightSemimodule R m₂ ℓm₂} + {M₃ : RawRightSemimodule R m₃ ℓm₃} + (open RawRightSemimodule) + (≈ᴹ₃-trans : Transitive (_≈ᴹ_ M₃)) {f : Carrierᴹ M₁ → Carrierᴹ M₂} {g : Carrierᴹ M₂ → Carrierᴹ M₃} where @@ -102,8 +105,8 @@ module _ IsRightSemimoduleHomomorphism M₂ M₃ g → IsRightSemimoduleHomomorphism M₁ M₃ (g ∘ f) isRightSemimoduleHomomorphism f-homo g-homo = record - { +ᴹ-isMonoidHomomorphism = isMonoidHomomorphism (≈ᴹ-trans M₃) F.+ᴹ-isMonoidHomomorphism G.+ᴹ-isMonoidHomomorphism - ; *ᵣ-homo = λ r x → ≈ᴹ-trans M₃ (G.⟦⟧-cong (F.*ᵣ-homo r x)) (G.*ᵣ-homo r (f x)) + { +ᴹ-isMonoidHomomorphism = isMonoidHomomorphism ≈ᴹ₃-trans F.+ᴹ-isMonoidHomomorphism G.+ᴹ-isMonoidHomomorphism + ; *ᵣ-homo = λ r x → ≈ᴹ₃-trans (G.⟦⟧-cong (F.*ᵣ-homo r x)) (G.*ᵣ-homo r (f x)) } where module F = IsRightSemimoduleHomomorphism f-homo; module G = IsRightSemimoduleHomomorphism g-homo isRightSemimoduleMonomorphism : IsRightSemimoduleMonomorphism M₁ M₂ f → @@ -123,11 +126,12 @@ module _ } where module F = IsRightSemimoduleIsomorphism f-iso; module G = IsRightSemimoduleIsomorphism g-iso module _ - {ring : Ring r ℓr} - {M₁ : RightModule ring m₁ ℓm₁} - {M₂ : RightModule ring m₂ ℓm₂} - {M₃ : RightModule ring m₃ ℓm₃} - (open RightModule) + {R : Set r} + {M₁ : RawRightModule R m₁ ℓm₁} + {M₂ : RawRightModule R m₂ ℓm₂} + {M₃ : RawRightModule R m₃ ℓm₃} + (open RawRightModule) + (≈ᴹ₃-trans : Transitive (_≈ᴹ_ M₃)) {f : Carrierᴹ M₁ → Carrierᴹ M₂} {g : Carrierᴹ M₂ → Carrierᴹ M₃} where @@ -136,8 +140,8 @@ module _ IsRightModuleHomomorphism M₂ M₃ g → IsRightModuleHomomorphism M₁ M₃ (g ∘ f) isRightModuleHomomorphism f-homo g-homo = record - { +ᴹ-isGroupHomomorphism = isGroupHomomorphism (≈ᴹ-trans M₃) F.+ᴹ-isGroupHomomorphism G.+ᴹ-isGroupHomomorphism - ; *ᵣ-homo = λ r x → ≈ᴹ-trans M₃ (G.⟦⟧-cong (F.*ᵣ-homo r x)) (G.*ᵣ-homo r (f x)) + { +ᴹ-isGroupHomomorphism = isGroupHomomorphism ≈ᴹ₃-trans F.+ᴹ-isGroupHomomorphism G.+ᴹ-isGroupHomomorphism + ; *ᵣ-homo = λ r x → ≈ᴹ₃-trans (G.⟦⟧-cong (F.*ᵣ-homo r x)) (G.*ᵣ-homo r (f x)) } where module F = IsRightModuleHomomorphism f-homo; module G = IsRightModuleHomomorphism g-homo isRightModuleMonomorphism : IsRightModuleMonomorphism M₁ M₂ f → @@ -157,12 +161,13 @@ module _ } where module F = IsRightModuleIsomorphism f-iso; module G = IsRightModuleIsomorphism g-iso module _ - {R-semiring : Semiring r ℓr} - {S-semiring : Semiring s ℓs} - {M₁ : Bisemimodule R-semiring S-semiring m₁ ℓm₁} - {M₂ : Bisemimodule R-semiring S-semiring m₂ ℓm₂} - {M₃ : Bisemimodule R-semiring S-semiring m₃ ℓm₃} - (open Bisemimodule) + {R : Set r} + {S : Set s} + {M₁ : RawBisemimodule R S m₁ ℓm₁} + {M₂ : RawBisemimodule R S m₂ ℓm₂} + {M₃ : RawBisemimodule R S m₃ ℓm₃} + (open RawBisemimodule) + (≈ᴹ₃-trans : Transitive (_≈ᴹ_ M₃)) {f : Carrierᴹ M₁ → Carrierᴹ M₂} {g : Carrierᴹ M₂ → Carrierᴹ M₃} where @@ -171,9 +176,9 @@ module _ IsBisemimoduleHomomorphism M₂ M₃ g → IsBisemimoduleHomomorphism M₁ M₃ (g ∘ f) isBisemimoduleHomomorphism f-homo g-homo = record - { +ᴹ-isMonoidHomomorphism = isMonoidHomomorphism (≈ᴹ-trans M₃) F.+ᴹ-isMonoidHomomorphism G.+ᴹ-isMonoidHomomorphism - ; *ₗ-homo = λ r x → ≈ᴹ-trans M₃ (G.⟦⟧-cong (F.*ₗ-homo r x)) (G.*ₗ-homo r (f x)) - ; *ᵣ-homo = λ r x → ≈ᴹ-trans M₃ (G.⟦⟧-cong (F.*ᵣ-homo r x)) (G.*ᵣ-homo r (f x)) + { +ᴹ-isMonoidHomomorphism = isMonoidHomomorphism ≈ᴹ₃-trans F.+ᴹ-isMonoidHomomorphism G.+ᴹ-isMonoidHomomorphism + ; *ₗ-homo = λ r x → ≈ᴹ₃-trans (G.⟦⟧-cong (F.*ₗ-homo r x)) (G.*ₗ-homo r (f x)) + ; *ᵣ-homo = λ r x → ≈ᴹ₃-trans (G.⟦⟧-cong (F.*ᵣ-homo r x)) (G.*ᵣ-homo r (f x)) } where module F = IsBisemimoduleHomomorphism f-homo; module G = IsBisemimoduleHomomorphism g-homo isBisemimoduleMonomorphism : IsBisemimoduleMonomorphism M₁ M₂ f → @@ -193,12 +198,13 @@ module _ } where module F = IsBisemimoduleIsomorphism f-iso; module G = IsBisemimoduleIsomorphism g-iso module _ - {R-ring : Ring r ℓr} - {S-ring : Ring s ℓs} - {M₁ : Bimodule R-ring S-ring m₁ ℓm₁} - {M₂ : Bimodule R-ring S-ring m₂ ℓm₂} - {M₃ : Bimodule R-ring S-ring m₃ ℓm₃} - (open Bimodule) + {R : Set r} + {S : Set s} + {M₁ : RawBimodule R S m₁ ℓm₁} + {M₂ : RawBimodule R S m₂ ℓm₂} + {M₃ : RawBimodule R S m₃ ℓm₃} + (open RawBimodule) + (≈ᴹ₃-trans : Transitive (_≈ᴹ_ M₃)) {f : Carrierᴹ M₁ → Carrierᴹ M₂} {g : Carrierᴹ M₂ → Carrierᴹ M₃} where @@ -207,9 +213,9 @@ module _ IsBimoduleHomomorphism M₂ M₃ g → IsBimoduleHomomorphism M₁ M₃ (g ∘ f) isBimoduleHomomorphism f-homo g-homo = record - { +ᴹ-isGroupHomomorphism = isGroupHomomorphism (≈ᴹ-trans M₃) F.+ᴹ-isGroupHomomorphism G.+ᴹ-isGroupHomomorphism - ; *ₗ-homo = λ r x → ≈ᴹ-trans M₃ (G.⟦⟧-cong (F.*ₗ-homo r x)) (G.*ₗ-homo r (f x)) - ; *ᵣ-homo = λ r x → ≈ᴹ-trans M₃ (G.⟦⟧-cong (F.*ᵣ-homo r x)) (G.*ᵣ-homo r (f x)) + { +ᴹ-isGroupHomomorphism = isGroupHomomorphism ≈ᴹ₃-trans F.+ᴹ-isGroupHomomorphism G.+ᴹ-isGroupHomomorphism + ; *ₗ-homo = λ r x → ≈ᴹ₃-trans (G.⟦⟧-cong (F.*ₗ-homo r x)) (G.*ₗ-homo r (f x)) + ; *ᵣ-homo = λ r x → ≈ᴹ₃-trans (G.⟦⟧-cong (F.*ᵣ-homo r x)) (G.*ᵣ-homo r (f x)) } where module F = IsBimoduleHomomorphism f-homo; module G = IsBimoduleHomomorphism g-homo isBimoduleMonomorphism : IsBimoduleMonomorphism M₁ M₂ f → @@ -229,11 +235,12 @@ module _ } where module F = IsBimoduleIsomorphism f-iso; module G = IsBimoduleIsomorphism g-iso module _ - {commutativeSemiring : CommutativeSemiring r ℓr} - {M₁ : Semimodule commutativeSemiring m₁ ℓm₁} - {M₂ : Semimodule commutativeSemiring m₂ ℓm₂} - {M₃ : Semimodule commutativeSemiring m₃ ℓm₃} - (open Semimodule) + {R : Set r} + {M₁ : RawSemimodule R m₁ ℓm₁} + {M₂ : RawSemimodule R m₂ ℓm₂} + {M₃ : RawSemimodule R m₃ ℓm₃} + (open RawSemimodule) + (≈ᴹ₃-trans : Transitive (_≈ᴹ_ M₃)) {f : Carrierᴹ M₁ → Carrierᴹ M₂} {g : Carrierᴹ M₂ → Carrierᴹ M₃} where @@ -242,7 +249,7 @@ module _ IsSemimoduleHomomorphism M₂ M₃ g → IsSemimoduleHomomorphism M₁ M₃ (g ∘ f) isSemimoduleHomomorphism f-homo g-homo = record - { isBisemimoduleHomomorphism = isBisemimoduleHomomorphism F.isBisemimoduleHomomorphism G.isBisemimoduleHomomorphism + { isBisemimoduleHomomorphism = isBisemimoduleHomomorphism ≈ᴹ₃-trans F.isBisemimoduleHomomorphism G.isBisemimoduleHomomorphism } where module F = IsSemimoduleHomomorphism f-homo; module G = IsSemimoduleHomomorphism g-homo isSemimoduleMonomorphism : IsSemimoduleMonomorphism M₁ M₂ f → @@ -262,11 +269,12 @@ module _ } where module F = IsSemimoduleIsomorphism f-iso; module G = IsSemimoduleIsomorphism g-iso module _ - {commutativeRing : CommutativeRing r ℓr} - {M₁ : Module commutativeRing m₁ ℓm₁} - {M₂ : Module commutativeRing m₂ ℓm₂} - {M₃ : Module commutativeRing m₃ ℓm₃} - (open Module) + {R : Set r} + {M₁ : RawModule R m₁ ℓm₁} + {M₂ : RawModule R m₂ ℓm₂} + {M₃ : RawModule R m₃ ℓm₃} + (open RawModule) + (≈ᴹ₃-trans : Transitive (_≈ᴹ_ M₃)) {f : Carrierᴹ M₁ → Carrierᴹ M₂} {g : Carrierᴹ M₂ → Carrierᴹ M₃} where @@ -275,7 +283,7 @@ module _ IsModuleHomomorphism M₂ M₃ g → IsModuleHomomorphism M₁ M₃ (g ∘ f) isModuleHomomorphism f-homo g-homo = record - { isBimoduleHomomorphism = isBimoduleHomomorphism F.isBimoduleHomomorphism G.isBimoduleHomomorphism + { isBimoduleHomomorphism = isBimoduleHomomorphism ≈ᴹ₃-trans F.isBimoduleHomomorphism G.isBimoduleHomomorphism } where module F = IsModuleHomomorphism f-homo; module G = IsModuleHomomorphism g-homo isModuleMonomorphism : IsModuleMonomorphism M₁ M₂ f → diff --git a/src/Algebra/Module/Morphism/Construct/Identity.agda b/src/Algebra/Module/Morphism/Construct/Identity.agda index 55735c2b24..54a08af98c 100644 --- a/src/Algebra/Module/Morphism/Construct/Identity.agda +++ b/src/Algebra/Module/Morphism/Construct/Identity.agda @@ -8,8 +8,7 @@ module Algebra.Module.Morphism.Construct.Identity where -open import Algebra.Bundles -open import Algebra.Module.Bundles +open import Algebra.Module.Bundles.Raw open import Algebra.Module.Morphism.Structures using ( module LeftSemimoduleMorphisms ; module LeftModuleMorphisms @@ -25,13 +24,13 @@ open import Data.Product.Base using (_,_) open import Function.Base using (id) import Function.Construct.Identity as Id open import Level using (Level) +open import Relation.Binary.Definitions using (Reflexive) private variable - r ℓr s ℓs m ℓm : Level + r s m ℓm : Level -module _ {semiring : Semiring r ℓr} (M : LeftSemimodule semiring m ℓm) where - open LeftSemimodule M using (≈ᴹ-refl) +module _ {R : Set r} (M : RawLeftSemimodule R m ℓm) (open RawLeftSemimodule M) (≈ᴹ-refl : Reflexive _≈ᴹ_) where open LeftSemimoduleMorphisms M M isLeftSemimoduleHomomorphism : IsLeftSemimoduleHomomorphism id @@ -52,8 +51,7 @@ module _ {semiring : Semiring r ℓr} (M : LeftSemimodule semiring m ℓm) where ; surjective = Id.surjective _ } -module _ {ring : Ring r ℓr} (M : LeftModule ring m ℓm) where - open LeftModule M using (≈ᴹ-refl) +module _ {R : Set r} (M : RawLeftModule R m ℓm) (open RawLeftModule M) (≈ᴹ-refl : Reflexive _≈ᴹ_) where open LeftModuleMorphisms M M isLeftModuleHomomorphism : IsLeftModuleHomomorphism id @@ -74,8 +72,7 @@ module _ {ring : Ring r ℓr} (M : LeftModule ring m ℓm) where ; surjective = Id.surjective _ } -module _ {semiring : Semiring r ℓr} (M : RightSemimodule semiring m ℓm) where - open RightSemimodule M using (≈ᴹ-refl) +module _ {R : Set r} (M : RawRightSemimodule R m ℓm) (open RawRightSemimodule M) (≈ᴹ-refl : Reflexive _≈ᴹ_) where open RightSemimoduleMorphisms M M isRightSemimoduleHomomorphism : IsRightSemimoduleHomomorphism id @@ -96,8 +93,7 @@ module _ {semiring : Semiring r ℓr} (M : RightSemimodule semiring m ℓm) wher ; surjective = Id.surjective _ } -module _ {ring : Ring r ℓr} (M : RightModule ring m ℓm) where - open RightModule M using (≈ᴹ-refl) +module _ {R : Set r} (M : RawRightModule R m ℓm) (open RawRightModule M) (≈ᴹ-refl : Reflexive _≈ᴹ_) where open RightModuleMorphisms M M isRightModuleHomomorphism : IsRightModuleHomomorphism id @@ -118,8 +114,7 @@ module _ {ring : Ring r ℓr} (M : RightModule ring m ℓm) where ; surjective = Id.surjective _ } -module _ {R-semiring : Semiring r ℓr} {S-semiring : Semiring s ℓs} (M : Bisemimodule R-semiring S-semiring m ℓm) where - open Bisemimodule M using (≈ᴹ-refl) +module _ {R : Set r} {S : Set s} (M : RawBisemimodule R S m ℓm) (open RawBisemimodule M) (≈ᴹ-refl : Reflexive _≈ᴹ_) where open BisemimoduleMorphisms M M isBisemimoduleHomomorphism : IsBisemimoduleHomomorphism id @@ -141,8 +136,7 @@ module _ {R-semiring : Semiring r ℓr} {S-semiring : Semiring s ℓs} (M : Bise ; surjective = Id.surjective _ } -module _ {R-ring : Ring r ℓr} {S-ring : Ring r ℓr} (M : Bimodule R-ring S-ring m ℓm) where - open Bimodule M using (≈ᴹ-refl) +module _ {R : Set r} {S : Set s} (M : RawBimodule R S m ℓm) (open RawBimodule M) (≈ᴹ-refl : Reflexive _≈ᴹ_) where open BimoduleMorphisms M M isBimoduleHomomorphism : IsBimoduleHomomorphism id @@ -164,13 +158,12 @@ module _ {R-ring : Ring r ℓr} {S-ring : Ring r ℓr} (M : Bimodule R-ring S-ri ; surjective = Id.surjective _ } -module _ {commutativeSemiring : CommutativeSemiring r ℓr} (M : Semimodule commutativeSemiring m ℓm) where - open Semimodule M using (≈ᴹ-refl) +module _ {R : Set r} (M : RawSemimodule R m ℓm) (open RawSemimodule M) (≈ᴹ-refl : Reflexive _≈ᴹ_) where open SemimoduleMorphisms M M isSemimoduleHomomorphism : IsSemimoduleHomomorphism id isSemimoduleHomomorphism = record - { isBisemimoduleHomomorphism = isBisemimoduleHomomorphism _ + { isBisemimoduleHomomorphism = isBisemimoduleHomomorphism _ ≈ᴹ-refl } isSemimoduleMonomorphism : IsSemimoduleMonomorphism id @@ -185,13 +178,12 @@ module _ {commutativeSemiring : CommutativeSemiring r ℓr} (M : Semimodule comm ; surjective = Id.surjective _ } -module _ {commutativeRing : CommutativeRing r ℓr} (M : Module commutativeRing m ℓm) where - open Module M using (≈ᴹ-refl) +module _ {R : Set r} (M : RawModule R m ℓm) (open RawModule M) (≈ᴹ-refl : Reflexive _≈ᴹ_) where open ModuleMorphisms M M isModuleHomomorphism : IsModuleHomomorphism id isModuleHomomorphism = record - { isBimoduleHomomorphism = isBimoduleHomomorphism _ + { isBimoduleHomomorphism = isBimoduleHomomorphism _ ≈ᴹ-refl } isModuleMonomorphism : IsModuleMonomorphism id diff --git a/src/Algebra/Module/Morphism/ModuleHomomorphism.agda b/src/Algebra/Module/Morphism/ModuleHomomorphism.agda index 704ec3863c..6bd7f2b4e6 100644 --- a/src/Algebra/Module/Morphism/ModuleHomomorphism.agda +++ b/src/Algebra/Module/Morphism/ModuleHomomorphism.agda @@ -17,10 +17,10 @@ module Algebra.Module.Morphism.ModuleHomomorphism {r ℓr m ℓm : Level} {ring : CommutativeRing r ℓr} (modA modB : Module ring m ℓm) - (open Module modA using () renaming (Carrierᴹ to A)) - (open Module modB using () renaming (Carrierᴹ to B)) + (open Module modA using () renaming (Carrierᴹ to A; rawModule to rawModA)) + (open Module modB using () renaming (Carrierᴹ to B; rawModule to rawModB)) {f : A → B} - (open MorphismStructures.ModuleMorphisms modA modB) + (open MorphismStructures.ModuleMorphisms rawModA rawModB) (isModuleHomomorphism : IsModuleHomomorphism f) where diff --git a/src/Algebra/Module/Morphism/Structures.agda b/src/Algebra/Module/Morphism/Structures.agda index 96fb5ffa0a..e874f4d341 100644 --- a/src/Algebra/Module/Morphism/Structures.agda +++ b/src/Algebra/Module/Morphism/Structures.agda @@ -7,8 +7,7 @@ module Algebra.Module.Morphism.Structures where -open import Algebra.Bundles -open import Algebra.Module.Bundles +open import Algebra.Module.Bundles.Raw import Algebra.Module.Morphism.Definitions as MorphismDefinitions import Algebra.Morphism.Structures as MorphismStructures open import Function.Definitions @@ -16,19 +15,18 @@ open import Level private variable - r ℓr s ℓs m₁ m₂ ℓm₁ ℓm₂ : Level + r s m₁ m₂ ℓm₁ ℓm₂ : Level module LeftSemimoduleMorphisms - {semiring : Semiring r ℓr} - (M₁ : LeftSemimodule semiring m₁ ℓm₁) - (M₂ : LeftSemimodule semiring m₂ ℓm₂) + {R : Set r} + (M₁ : RawLeftSemimodule R m₁ ℓm₁) + (M₂ : RawLeftSemimodule R m₂ ℓm₂) where - open Semiring semiring renaming (Carrier to R) - open LeftSemimodule M₁ renaming (Carrierᴹ to A; _*ₗ_ to _*ₗ₁_; _≈ᴹ_ to _≈ᴹ₁_) - open LeftSemimodule M₂ renaming (Carrierᴹ to B; _*ₗ_ to _*ₗ₂_; _≈ᴹ_ to _≈ᴹ₂_) + open RawLeftSemimodule M₁ renaming (Carrierᴹ to A; _*ₗ_ to _*ₗ₁_; _≈ᴹ_ to _≈ᴹ₁_) + open RawLeftSemimodule M₂ renaming (Carrierᴹ to B; _*ₗ_ to _*ₗ₂_; _≈ᴹ_ to _≈ᴹ₂_) open MorphismDefinitions R A B _≈ᴹ₂_ - open MorphismStructures.MonoidMorphisms (LeftSemimodule.+ᴹ-rawMonoid M₁) (LeftSemimodule.+ᴹ-rawMonoid M₂) + open MorphismStructures.MonoidMorphisms (RawLeftSemimodule.+ᴹ-rawMonoid M₁) (RawLeftSemimodule.+ᴹ-rawMonoid M₂) record IsLeftSemimoduleHomomorphism (⟦_⟧ : A → B) : Set (r ⊔ m₁ ⊔ ℓm₁ ⊔ ℓm₂) where field @@ -74,17 +72,16 @@ module LeftSemimoduleMorphisms renaming (isMagmaIsomorphism to +ᴹ-isMagmaIsomorphism) module LeftModuleMorphisms - {ring : Ring r ℓr} - (M₁ : LeftModule ring m₁ ℓm₁) - (M₂ : LeftModule ring m₂ ℓm₂) + {R : Set r} + (M₁ : RawLeftModule R m₁ ℓm₁) + (M₂ : RawLeftModule R m₂ ℓm₂) where - open Ring ring renaming (Carrier to R) - open LeftModule M₁ renaming (Carrierᴹ to A; _*ₗ_ to _*ₗ₁_; _≈ᴹ_ to _≈ᴹ₁_) - open LeftModule M₂ renaming (Carrierᴹ to B; _*ₗ_ to _*ₗ₂_; _≈ᴹ_ to _≈ᴹ₂_) + open RawLeftModule M₁ renaming (Carrierᴹ to A; _*ₗ_ to _*ₗ₁_; _≈ᴹ_ to _≈ᴹ₁_) + open RawLeftModule M₂ renaming (Carrierᴹ to B; _*ₗ_ to _*ₗ₂_; _≈ᴹ_ to _≈ᴹ₂_) open MorphismDefinitions R A B _≈ᴹ₂_ - open MorphismStructures.GroupMorphisms (LeftModule.+ᴹ-rawGroup M₁) (LeftModule.+ᴹ-rawGroup M₂) - open LeftSemimoduleMorphisms (LeftModule.leftSemimodule M₁) (LeftModule.leftSemimodule M₂) + open MorphismStructures.GroupMorphisms (RawLeftModule.+ᴹ-rawGroup M₁) (RawLeftModule.+ᴹ-rawGroup M₂) + open LeftSemimoduleMorphisms (RawLeftModule.rawLeftSemimodule M₁) (RawLeftModule.rawLeftSemimodule M₂) record IsLeftModuleHomomorphism (⟦_⟧ : A → B) : Set (r ⊔ m₁ ⊔ ℓm₁ ⊔ ℓm₂) where field @@ -148,16 +145,15 @@ module LeftModuleMorphisms } module RightSemimoduleMorphisms - {semiring : Semiring r ℓr} - (M₁ : RightSemimodule semiring m₁ ℓm₁) - (M₂ : RightSemimodule semiring m₂ ℓm₂) + {R : Set r} + (M₁ : RawRightSemimodule R m₁ ℓm₁) + (M₂ : RawRightSemimodule R m₂ ℓm₂) where - open Semiring semiring renaming (Carrier to R) - open RightSemimodule M₁ renaming (Carrierᴹ to A; _*ᵣ_ to _*ᵣ₁_; _≈ᴹ_ to _≈ᴹ₁_) - open RightSemimodule M₂ renaming (Carrierᴹ to B; _*ᵣ_ to _*ᵣ₂_; _≈ᴹ_ to _≈ᴹ₂_) + open RawRightSemimodule M₁ renaming (Carrierᴹ to A; _*ᵣ_ to _*ᵣ₁_; _≈ᴹ_ to _≈ᴹ₁_) + open RawRightSemimodule M₂ renaming (Carrierᴹ to B; _*ᵣ_ to _*ᵣ₂_; _≈ᴹ_ to _≈ᴹ₂_) open MorphismDefinitions R A B _≈ᴹ₂_ - open MorphismStructures.MonoidMorphisms (RightSemimodule.+ᴹ-rawMonoid M₁) (RightSemimodule.+ᴹ-rawMonoid M₂) + open MorphismStructures.MonoidMorphisms (RawRightSemimodule.+ᴹ-rawMonoid M₁) (RawRightSemimodule.+ᴹ-rawMonoid M₂) record IsRightSemimoduleHomomorphism (⟦_⟧ : A → B) : Set (r ⊔ m₁ ⊔ ℓm₁ ⊔ ℓm₂) where field @@ -203,17 +199,16 @@ module RightSemimoduleMorphisms renaming (isMagmaIsomorphism to +ᴹ-isMagmaIsomorphism) module RightModuleMorphisms - {ring : Ring r ℓr} - (M₁ : RightModule ring m₁ ℓm₁) - (M₂ : RightModule ring m₂ ℓm₂) + {R : Set r} + (M₁ : RawRightModule R m₁ ℓm₁) + (M₂ : RawRightModule R m₂ ℓm₂) where - open Ring ring renaming (Carrier to R) - open RightModule M₁ renaming (Carrierᴹ to A; _*ᵣ_ to _*ᵣ₁_; _≈ᴹ_ to _≈ᴹ₁_) - open RightModule M₂ renaming (Carrierᴹ to B; _*ᵣ_ to _*ᵣ₂_; _≈ᴹ_ to _≈ᴹ₂_) + open RawRightModule M₁ renaming (Carrierᴹ to A; _*ᵣ_ to _*ᵣ₁_; _≈ᴹ_ to _≈ᴹ₁_) + open RawRightModule M₂ renaming (Carrierᴹ to B; _*ᵣ_ to _*ᵣ₂_; _≈ᴹ_ to _≈ᴹ₂_) open MorphismDefinitions R A B _≈ᴹ₂_ - open MorphismStructures.GroupMorphisms (RightModule.+ᴹ-rawGroup M₁) (RightModule.+ᴹ-rawGroup M₂) - open RightSemimoduleMorphisms (RightModule.rightSemimodule M₁) (RightModule.rightSemimodule M₂) + open MorphismStructures.GroupMorphisms (RawRightModule.+ᴹ-rawGroup M₁) (RawRightModule.+ᴹ-rawGroup M₂) + open RightSemimoduleMorphisms (RawRightModule.rawRightSemimodule M₁) (RawRightModule.rawRightSemimodule M₂) record IsRightModuleHomomorphism (⟦_⟧ : A → B) : Set (r ⊔ m₁ ⊔ ℓm₁ ⊔ ℓm₂) where field @@ -276,21 +271,19 @@ module RightModuleMorphisms } module BisemimoduleMorphisms - {R-semiring : Semiring r ℓr} - {S-semiring : Semiring s ℓs} - (M₁ : Bisemimodule R-semiring S-semiring m₁ ℓm₁) - (M₂ : Bisemimodule R-semiring S-semiring m₂ ℓm₂) + {R : Set r} + {S : Set s} + (M₁ : RawBisemimodule R S m₁ ℓm₁) + (M₂ : RawBisemimodule R S m₂ ℓm₂) where - open Semiring R-semiring renaming (Carrier to R) - open Semiring S-semiring renaming (Carrier to S) - open Bisemimodule M₁ renaming (Carrierᴹ to A; _*ₗ_ to _*ₗ₁_; _*ᵣ_ to _*ᵣ₁_; _≈ᴹ_ to _≈ᴹ₁_) - open Bisemimodule M₂ renaming (Carrierᴹ to B; _*ₗ_ to _*ₗ₂_; _*ᵣ_ to _*ᵣ₂_; _≈ᴹ_ to _≈ᴹ₂_) + open RawBisemimodule M₁ renaming (Carrierᴹ to A; _*ₗ_ to _*ₗ₁_; _*ᵣ_ to _*ᵣ₁_; _≈ᴹ_ to _≈ᴹ₁_) + open RawBisemimodule M₂ renaming (Carrierᴹ to B; _*ₗ_ to _*ₗ₂_; _*ᵣ_ to _*ᵣ₂_; _≈ᴹ_ to _≈ᴹ₂_) open MorphismDefinitions R A B _≈ᴹ₂_ using (Homomorphicₗ) open MorphismDefinitions S A B _≈ᴹ₂_ using (Homomorphicᵣ) - open MorphismStructures.MonoidMorphisms (Bisemimodule.+ᴹ-rawMonoid M₁) (Bisemimodule.+ᴹ-rawMonoid M₂) - open LeftSemimoduleMorphisms (Bisemimodule.leftSemimodule M₁) (Bisemimodule.leftSemimodule M₂) - open RightSemimoduleMorphisms (Bisemimodule.rightSemimodule M₁) (Bisemimodule.rightSemimodule M₂) + open MorphismStructures.MonoidMorphisms (RawBisemimodule.+ᴹ-rawMonoid M₁) (RawBisemimodule.+ᴹ-rawMonoid M₂) + open LeftSemimoduleMorphisms (RawBisemimodule.rawLeftSemimodule M₁) (RawBisemimodule.rawLeftSemimodule M₂) + open RightSemimoduleMorphisms (RawBisemimodule.rawRightSemimodule M₁) (RawBisemimodule.rawRightSemimodule M₂) record IsBisemimoduleHomomorphism (⟦_⟧ : A → B) : Set (r ⊔ s ⊔ m₁ ⊔ ℓm₁ ⊔ ℓm₂) where field @@ -358,22 +351,20 @@ module BisemimoduleMorphisms } module BimoduleMorphisms - {R-ring : Ring r ℓr} - {S-ring : Ring s ℓs} - (M₁ : Bimodule R-ring S-ring m₁ ℓm₁) - (M₂ : Bimodule R-ring S-ring m₂ ℓm₂) + {R : Set r} + {S : Set s} + (M₁ : RawBimodule R S m₁ ℓm₁) + (M₂ : RawBimodule R S m₂ ℓm₂) where - open Ring R-ring renaming (Carrier to R) - open Ring S-ring renaming (Carrier to S) - open Bimodule M₁ renaming (Carrierᴹ to A; _*ₗ_ to _*ₗ₁_; _*ᵣ_ to _*ᵣ₁_; _≈ᴹ_ to _≈ᴹ₁_) - open Bimodule M₂ renaming (Carrierᴹ to B; _*ₗ_ to _*ₗ₂_; _*ᵣ_ to _*ᵣ₂_; _≈ᴹ_ to _≈ᴹ₂_) + open RawBimodule M₁ renaming (Carrierᴹ to A; _*ₗ_ to _*ₗ₁_; _*ᵣ_ to _*ᵣ₁_; _≈ᴹ_ to _≈ᴹ₁_) + open RawBimodule M₂ renaming (Carrierᴹ to B; _*ₗ_ to _*ₗ₂_; _*ᵣ_ to _*ᵣ₂_; _≈ᴹ_ to _≈ᴹ₂_) open MorphismDefinitions R A B _≈ᴹ₂_ using (Homomorphicₗ) open MorphismDefinitions S A B _≈ᴹ₂_ using (Homomorphicᵣ) - open MorphismStructures.GroupMorphisms (Bimodule.+ᴹ-rawGroup M₁) (Bimodule.+ᴹ-rawGroup M₂) - open LeftModuleMorphisms (Bimodule.leftModule M₁) (Bimodule.leftModule M₂) - open RightModuleMorphisms (Bimodule.rightModule M₁) (Bimodule.rightModule M₂) - open BisemimoduleMorphisms (Bimodule.bisemimodule M₁) (Bimodule.bisemimodule M₂) + open MorphismStructures.GroupMorphisms (RawBimodule.+ᴹ-rawGroup M₁) (RawBimodule.+ᴹ-rawGroup M₂) + open LeftModuleMorphisms (RawBimodule.rawLeftModule M₁) (RawBimodule.rawLeftModule M₂) + open RightModuleMorphisms (RawBimodule.rawRightModule M₁) (RawBimodule.rawRightModule M₂) + open BisemimoduleMorphisms (RawBimodule.rawBisemimodule M₁) (RawBimodule.rawBisemimodule M₂) record IsBimoduleHomomorphism (⟦_⟧ : A → B) : Set (r ⊔ s ⊔ m₁ ⊔ ℓm₁ ⊔ ℓm₂) where field @@ -492,14 +483,14 @@ module BimoduleMorphisms } module SemimoduleMorphisms - {commutativeSemiring : CommutativeSemiring r ℓr} - (M₁ : Semimodule commutativeSemiring m₁ ℓm₁) - (M₂ : Semimodule commutativeSemiring m₂ ℓm₂) + {R : Set r} + (M₁ : RawSemimodule R m₁ ℓm₁) + (M₂ : RawSemimodule R m₂ ℓm₂) where - open Semimodule M₁ renaming (Carrierᴹ to A; _≈ᴹ_ to _≈ᴹ₁_) - open Semimodule M₂ renaming (Carrierᴹ to B; _≈ᴹ_ to _≈ᴹ₂_) - open BisemimoduleMorphisms (Semimodule.bisemimodule M₁) (Semimodule.bisemimodule M₂) + open RawSemimodule M₁ renaming (Carrierᴹ to A; _≈ᴹ_ to _≈ᴹ₁_) + open RawSemimodule M₂ renaming (Carrierᴹ to B; _≈ᴹ_ to _≈ᴹ₂_) + open BisemimoduleMorphisms (RawSemimodule.rawBisemimodule M₁) (RawSemimodule.rawBisemimodule M₂) record IsSemimoduleHomomorphism (⟦_⟧ : A → B) : Set (r ⊔ m₁ ⊔ ℓm₁ ⊔ ℓm₂) where field @@ -544,15 +535,15 @@ module SemimoduleMorphisms ) module ModuleMorphisms - {commutativeRing : CommutativeRing r ℓr} - (M₁ : Module commutativeRing m₁ ℓm₁) - (M₂ : Module commutativeRing m₂ ℓm₂) + {R : Set r} + (M₁ : RawModule R m₁ ℓm₁) + (M₂ : RawModule R m₂ ℓm₂) where - open Module M₁ renaming (Carrierᴹ to A; _≈ᴹ_ to _≈ᴹ₁_) - open Module M₂ renaming (Carrierᴹ to B; _≈ᴹ_ to _≈ᴹ₂_) - open BimoduleMorphisms (Module.bimodule M₁) (Module.bimodule M₂) - open SemimoduleMorphisms (Module.semimodule M₁) (Module.semimodule M₂) + open RawModule M₁ renaming (Carrierᴹ to A; _≈ᴹ_ to _≈ᴹ₁_) + open RawModule M₂ renaming (Carrierᴹ to B; _≈ᴹ_ to _≈ᴹ₂_) + open BimoduleMorphisms (RawModule.rawBimodule M₁) (RawModule.rawBimodule M₂) + open SemimoduleMorphisms (RawModule.rawBisemimodule M₁) (RawModule.rawBisemimodule M₂) record IsModuleHomomorphism (⟦_⟧ : A → B) : Set (r ⊔ m₁ ⊔ ℓm₁ ⊔ ℓm₂) where field diff --git a/src/Algebra/Morphism.agda b/src/Algebra/Morphism.agda index 1bba00bab2..45b7a46a1d 100644 --- a/src/Algebra/Morphism.agda +++ b/src/Algebra/Morphism.agda @@ -14,7 +14,7 @@ import Algebra.Properties.Group as GroupP open import Function.Base open import Level open import Relation.Binary.Core using (Rel; _Preserves_⟶_) -import Relation.Binary.Reasoning.Setoid as EqR +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning private variable @@ -137,7 +137,7 @@ module _ {c₁ ℓ₁ c₂ ℓ₂} open IsMonoidMorphism mn-homo public ⁻¹-homo : Homomorphic₁ ⟦_⟧ F._⁻¹ T._⁻¹ - ⁻¹-homo x = let open EqR T.setoid in T.uniqueˡ-⁻¹ ⟦ x F.⁻¹ ⟧ ⟦ x ⟧ $ begin + ⁻¹-homo x = let open ≈-Reasoning T.setoid in T.uniqueˡ-⁻¹ ⟦ x F.⁻¹ ⟧ ⟦ x ⟧ $ begin ⟦ x F.⁻¹ ⟧ T.∙ ⟦ x ⟧ ≈⟨ T.sym (∙-homo (x F.⁻¹) x) ⟩ ⟦ x F.⁻¹ F.∙ x ⟧ ≈⟨ ⟦⟧-cong (F.inverseˡ x) ⟩ ⟦ F.ε ⟧ ≈⟨ ε-homo ⟩ diff --git a/src/Algebra/Morphism/Consequences.agda b/src/Algebra/Morphism/Consequences.agda index e297a8dbc9..b451a87088 100644 --- a/src/Algebra/Morphism/Consequences.agda +++ b/src/Algebra/Morphism/Consequences.agda @@ -13,7 +13,7 @@ open import Algebra.Morphism.Definitions open import Data.Product.Base using (_,_) open import Function.Base using (id; _∘_) open import Function.Definitions -import Relation.Binary.Reasoning.Setoid as EqR +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning ------------------------------------------------------------------------ -- If f and g are mutually inverse maps between A and B, g is congruent, @@ -34,7 +34,7 @@ module _ {α α= β β=} (M₁ : Magma α α=) (M₂ : Magma β β=) where g (f (g x) ∙₂ f (g y)) ≈⟨ g-cong (homo (g x) (g y)) ⟨ g (f (g x ∙₁ g y)) ≈⟨ invʳ M₂.refl ⟩ g x ∙₁ g y ∎ - where open EqR M₁.setoid + where open ≈-Reasoning M₁.setoid homomorphic₂-inj : ∀ {f g} → Injective _≈₁_ _≈₂_ f → Inverseˡ _≈₁_ _≈₂_ f g → @@ -45,4 +45,4 @@ module _ {α α= β β=} (M₁ : Magma α α=) (M₂ : Magma β β=) where x ∙₂ y ≈⟨ M₂.∙-cong (invˡ M₁.refl) (invˡ M₁.refl) ⟨ f (g x) ∙₂ f (g y) ≈⟨ homo (g x) (g y) ⟨ f (g x ∙₁ g y) ∎) - where open EqR M₂.setoid + where open ≈-Reasoning M₂.setoid diff --git a/src/Algebra/Morphism/GroupMonomorphism.agda b/src/Algebra/Morphism/GroupMonomorphism.agda index 51cdcbafe7..4b3349d141 100644 --- a/src/Algebra/Morphism/GroupMonomorphism.agda +++ b/src/Algebra/Morphism/GroupMonomorphism.agda @@ -27,7 +27,7 @@ open RawGroup G₂ renaming open import Algebra.Definitions open import Algebra.Structures open import Data.Product.Base using (_,_) -import Relation.Binary.Reasoning.Setoid as SetoidReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning ------------------------------------------------------------------------ -- Re-export all properties of monoid monomorphisms @@ -41,7 +41,7 @@ open import Algebra.Morphism.MonoidMonomorphism module _ (◦-isMagma : IsMagma _≈₂_ _◦_) where open IsMagma ◦-isMagma renaming (∙-cong to ◦-cong) - open SetoidReasoning setoid + open ≈-Reasoning setoid inverseˡ : LeftInverse _≈₂_ ε₂ _⁻¹₂ _◦_ → LeftInverse _≈₁_ ε₁ _⁻¹₁ _∙_ inverseˡ invˡ x = injective (begin @@ -72,7 +72,7 @@ module _ (◦-isMagma : IsMagma _≈₂_ _◦_) where module _ (◦-isAbelianGroup : IsAbelianGroup _≈₂_ _◦_ ε₂ _⁻¹₂) where open IsAbelianGroup ◦-isAbelianGroup renaming (∙-cong to ◦-cong; ⁻¹-cong to ⁻¹₂-cong) - open SetoidReasoning setoid + open ≈-Reasoning setoid ⁻¹-distrib-∙ : (∀ x y → (x ◦ y) ⁻¹₂ ≈₂ (x ⁻¹₂) ◦ (y ⁻¹₂)) → (∀ x y → (x ∙ y) ⁻¹₁ ≈₁ (x ⁻¹₁) ∙ (y ⁻¹₁)) ⁻¹-distrib-∙ ⁻¹-distrib-∙ x y = injective (begin diff --git a/src/Algebra/Morphism/MagmaMonomorphism.agda b/src/Algebra/Morphism/MagmaMonomorphism.agda index 87027ba5e1..fbe08a5861 100644 --- a/src/Algebra/Morphism/MagmaMonomorphism.agda +++ b/src/Algebra/Morphism/MagmaMonomorphism.agda @@ -27,7 +27,7 @@ open import Algebra.Structures open import Algebra.Definitions open import Data.Product.Base using (map) open import Data.Sum.Base using (inj₁; inj₂) -import Relation.Binary.Reasoning.Setoid as SetoidReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning import Relation.Binary.Morphism.RelMonomorphism isRelMonomorphism as RelMorphism ------------------------------------------------------------------------ @@ -36,7 +36,7 @@ import Relation.Binary.Morphism.RelMonomorphism isRelMonomorphism as RelMorphism module _ (◦-isMagma : IsMagma _≈₂_ _◦_) where open IsMagma ◦-isMagma renaming (∙-cong to ◦-cong) - open SetoidReasoning setoid + open ≈-Reasoning setoid cong : Congruent₂ _≈₁_ _∙_ cong {x} {y} {u} {v} x≈y u≈v = injective (begin diff --git a/src/Algebra/Morphism/MonoidMonomorphism.agda b/src/Algebra/Morphism/MonoidMonomorphism.agda index bfb412302e..1b918295af 100644 --- a/src/Algebra/Morphism/MonoidMonomorphism.agda +++ b/src/Algebra/Morphism/MonoidMonomorphism.agda @@ -25,7 +25,7 @@ open RawMonoid M₂ renaming (Carrier to B; _≈_ to _≈₂_; _∙_ to _◦_; open import Algebra.Definitions open import Algebra.Structures open import Data.Product.Base using (map) -import Relation.Binary.Reasoning.Setoid as SetoidReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning ------------------------------------------------------------------------ -- Re-export all properties of magma monomorphisms @@ -39,7 +39,7 @@ open import Algebra.Morphism.MagmaMonomorphism module _ (◦-isMagma : IsMagma _≈₂_ _◦_) where open IsMagma ◦-isMagma renaming (∙-cong to ◦-cong) - open SetoidReasoning setoid + open ≈-Reasoning setoid identityˡ : LeftIdentity _≈₂_ ε₂ _◦_ → LeftIdentity _≈₁_ ε₁ _∙_ identityˡ idˡ x = injective (begin diff --git a/src/Algebra/Morphism/RingMonomorphism.agda b/src/Algebra/Morphism/RingMonomorphism.agda index 35a191a0f8..40e7a80d83 100644 --- a/src/Algebra/Morphism/RingMonomorphism.agda +++ b/src/Algebra/Morphism/RingMonomorphism.agda @@ -29,7 +29,7 @@ open RawRing R₂ renaming open import Algebra.Definitions open import Algebra.Structures open import Data.Product.Base using (proj₁; proj₂; _,_) -import Relation.Binary.Reasoning.Setoid as SetoidReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning ------------------------------------------------------------------------ -- Re-export all properties of group and monoid monomorphisms @@ -83,7 +83,7 @@ module _ (+-isGroup : IsGroup _≈₂_ _⊕_ 0#₂ ⊝_) open IsGroup +-isGroup hiding (setoid; refl; sym) open IsMagma *-isMagma renaming (∙-cong to ◦-cong) - open SetoidReasoning setoid + open ≈-Reasoning setoid distribˡ : _DistributesOverˡ_ _≈₂_ _⊛_ _⊕_ → _DistributesOverˡ_ _≈₁_ _*_ _+_ distribˡ distribˡ x y z = injective (begin diff --git a/src/Algebra/Properties/CommutativeSemigroup/Divisibility.agda b/src/Algebra/Properties/CommutativeSemigroup/Divisibility.agda index da8d3abf45..782b094fbe 100644 --- a/src/Algebra/Properties/CommutativeSemigroup/Divisibility.agda +++ b/src/Algebra/Properties/CommutativeSemigroup/Divisibility.agda @@ -8,7 +8,7 @@ open import Algebra using (CommutativeSemigroup) open import Data.Product.Base using (_,_) -import Relation.Binary.Reasoning.Setoid as EqReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning module Algebra.Properties.CommutativeSemigroup.Divisibility {a ℓ} (CS : CommutativeSemigroup a ℓ) @@ -16,7 +16,7 @@ module Algebra.Properties.CommutativeSemigroup.Divisibility open CommutativeSemigroup CS open import Algebra.Properties.CommutativeSemigroup CS using (x∙yz≈xz∙y; x∙yz≈y∙xz) -open EqReasoning setoid +open ≈-Reasoning setoid ------------------------------------------------------------------------ -- Re-export the contents of divisibility over semigroups diff --git a/src/Algebra/Properties/Monoid/Mult.agda b/src/Algebra/Properties/Monoid/Mult.agda index 88c4258fcb..2963523b77 100644 --- a/src/Algebra/Properties/Monoid/Mult.agda +++ b/src/Algebra/Properties/Monoid/Mult.agda @@ -51,6 +51,12 @@ open import Algebra.Definitions.RawMonoid rawMonoid public -- _×_ is homomorphic with respect to _ℕ+_/_+_. +×-homo-0 : ∀ x → 0 × x ≈ 0# +×-homo-0 x = refl + +×-homo-1 : ∀ x → 1 × x ≈ x +×-homo-1 = +-identityʳ + ×-homo-+ : ∀ x m n → (m ℕ.+ n) × x ≈ m × x + n × x ×-homo-+ x 0 n = sym (+-identityˡ (n × x)) ×-homo-+ x (suc m) n = begin diff --git a/src/Algebra/Properties/Semiring/Binomial.agda b/src/Algebra/Properties/Semiring/Binomial.agda index 6731018af6..bbd9adce21 100644 --- a/src/Algebra/Properties/Semiring/Binomial.agda +++ b/src/Algebra/Properties/Semiring/Binomial.agda @@ -10,10 +10,10 @@ open import Algebra.Bundles using (Semiring) open import Data.Bool.Base using (true) -open import Data.Nat.Base as Nat hiding (_+_; _*_; _^_) +open import Data.Nat.Base as ℕ hiding (_+_; _*_; _^_) open import Data.Nat.Combinatorics using (_C_; nCn≡1; nC1≡n; nCk+nC[k+1]≡[n+1]C[k+1]) -open import Data.Nat.Properties as Nat +open import Data.Nat.Properties as ℕ using (<⇒<ᵇ; n<1+n; n∸n≡0; +-∸-assoc) open import Data.Fin.Base as Fin using (Fin; zero; suc; toℕ; fromℕ; inject₁) @@ -154,8 +154,8 @@ y*lemma x*y≈y*x {n} j = begin -- Now, a lemma characterising the sum of the term₁ and term₂ expressions private - n<ᵇ1+n : ∀ n → (n Nat.<ᵇ suc n) ≡ true - n<ᵇ1+n n with true ← n Nat.<ᵇ suc n | _ ← <⇒<ᵇ (n<1+n n) = ≡.refl + n<ᵇ1+n : ∀ n → (n ℕ.<ᵇ suc n) ≡ true + n<ᵇ1+n n with true ← n ℕ.<ᵇ suc n | _ ← <⇒<ᵇ (n<1+n n) = ≡.refl term₁+term₂≈term : x * y ≈ y * x → ∀ n i → term₁ n i + term₂ n i ≈ binomialTerm (suc n) i @@ -193,7 +193,7 @@ term₁+term₂≈term x*y≈y*x n (suc i) with view i ≈⟨ +-congˡ (×-congˡ nC[k+1]≡nC[j+1]) ⟨ (nCk × [x^k+1]*[y^n-k]) + (nC[k+1] × [x^k+1]*[y^n-k]) ≈⟨ ×-homo-+ [x^k+1]*[y^n-k] nCk nC[k+1] ⟨ - (nCk Nat.+ nC[k+1]) × [x^k+1]*[y^n-k] + (nCk ℕ.+ nC[k+1]) × [x^k+1]*[y^n-k] ≡⟨ cong (_× [x^k+1]*[y^n-k]) (nCk+nC[k+1]≡[n+1]C[k+1] n k) ⟩ ((suc n) C (suc k)) × [x^k+1]*[y^n-k] ≡⟨⟩ @@ -216,7 +216,7 @@ term₁+term₂≈term x*y≈y*x n (suc i) with view i theorem : x * y ≈ y * x → ∀ n → (x + y) ^ n ≈ binomialExpansion n theorem x*y≈y*x zero = begin (x + y) ^ 0 ≡⟨⟩ - 1# ≈⟨ 1×-identityʳ 1# ⟨ + 1# ≈⟨ ×-homo-1 1# ⟨ 1 × 1# ≈⟨ *-identityʳ (1 × 1#) ⟨ (1 × 1#) * 1# ≈⟨ ×-assoc-* 1 1# 1# ⟩ 1 × (1# * 1#) ≡⟨⟩ diff --git a/src/Algebra/Properties/Semiring/Mult.agda b/src/Algebra/Properties/Semiring/Mult.agda index 801e4f1966..6301d274c6 100644 --- a/src/Algebra/Properties/Semiring/Mult.agda +++ b/src/Algebra/Properties/Semiring/Mult.agda @@ -6,7 +6,7 @@ {-# OPTIONS --cubical-compatible --safe #-} -open import Algebra +open import Algebra.Bundles using (Semiring) open import Data.Nat.Base as ℕ using (zero; suc) module Algebra.Properties.Semiring.Mult @@ -14,6 +14,7 @@ module Algebra.Properties.Semiring.Mult open Semiring S renaming (zero to *-zero) open import Relation.Binary.Reasoning.Setoid setoid +open import Algebra.Definitions _≈_ using (_IdempotentOn_) ------------------------------------------------------------------------ -- Re-export definition from the monoid @@ -23,21 +24,15 @@ open import Algebra.Properties.Monoid.Mult +-monoid public ------------------------------------------------------------------------ -- Properties of _×_ --- (_× 1#) is homomorphic with respect to _ℕ.*_/_*_. +-- (0 ×_) is (0# *_) -×1-homo-* : ∀ m n → (m ℕ.* n) × 1# ≈ (m × 1#) * (n × 1#) -×1-homo-* 0 n = sym (zeroˡ (n × 1#)) -×1-homo-* (suc m) n = begin - (n ℕ.+ m ℕ.* n) × 1# ≈⟨ ×-homo-+ 1# n (m ℕ.* n) ⟩ - n × 1# + (m ℕ.* n) × 1# ≈⟨ +-congˡ (×1-homo-* m n) ⟩ - n × 1# + (m × 1#) * (n × 1#) ≈⟨ +-congʳ (*-identityˡ _) ⟨ - 1# * (n × 1#) + (m × 1#) * (n × 1#) ≈⟨ distribʳ (n × 1#) 1# (m × 1#) ⟨ - (1# + m × 1#) * (n × 1#) ∎ +×-homo-0# : ∀ x → 0 × x ≈ 0# * x +×-homo-0# x = sym (zeroˡ x) --- (1 ×_) is the identity +-- (1 ×_) is (1# *_) -1×-identityʳ : ∀ x → 1 × x ≈ x -1×-identityʳ = +-identityʳ +×-homo-1# : ∀ x → 1 × x ≈ 1# * x +×-homo-1# x = trans (×-homo-1 x) (sym (*-identityˡ x)) -- (n ×_) commutes with _*_ @@ -60,3 +55,32 @@ open import Algebra.Properties.Monoid.Mult +-monoid public x * y + (n × x) * y ≈⟨ +-congˡ (×-assoc-* n _ _) ⟩ x * y + n × (x * y) ≡⟨⟩ suc n × (x * y) ∎ + +-- (_× x) is homomorphic with respect to _ℕ.*_/_*_ for idempotent x. + +idem-×-homo-* : ∀ m n {x} → (_*_ IdempotentOn x) → (m × x) * (n × x) ≈ (m ℕ.* n) × x +idem-×-homo-* m n {x} idem = begin + (m × x) * (n × x) ≈⟨ ×-assoc-* m x (n × x) ⟩ + m × (x * (n × x)) ≈⟨ ×-congʳ m (×-comm-* n x x) ⟩ + m × (n × (x * x)) ≈⟨ ×-assocˡ _ m n ⟩ + (m ℕ.* n) × (x * x) ≈⟨ ×-congʳ (m ℕ.* n) idem ⟩ + (m ℕ.* n) × x ∎ + +-- (_× 1#) is homomorphic with respect to _ℕ.*_/_*_. + +×1-homo-* : ∀ m n → (m ℕ.* n) × 1# ≈ (m × 1#) * (n × 1#) +×1-homo-* m n = sym (idem-×-homo-* m n (*-identityʳ 1#)) + +------------------------------------------------------------------------ +-- DEPRECATED NAMES +------------------------------------------------------------------------ +-- Please use the new names as continuing support for the old names is +-- not guaranteed. + +-- Version 2.1 + +1×-identityʳ = ×-homo-1 +{-# WARNING_ON_USAGE 1×-identityʳ +"Warning: 1×-identityʳ was deprecated in v2.1. +Please use ×-homo-1 instead. " +#-} diff --git a/src/Algebra/Solver/CommutativeMonoid.agda b/src/Algebra/Solver/CommutativeMonoid.agda index 8062ff34df..f15f83824f 100644 --- a/src/Algebra/Solver/CommutativeMonoid.agda +++ b/src/Algebra/Solver/CommutativeMonoid.agda @@ -22,16 +22,16 @@ open import Data.Vec.Base using (Vec; []; _∷_; lookup; replicate) open import Function.Base using (_∘_) -import Relation.Binary.Reasoning.Setoid as EqReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning import Relation.Binary.Reflection as Reflection import Relation.Nullary.Decidable as Dec import Data.Vec.Relation.Binary.Pointwise.Inductive as Pointwise -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary.Decidable using (Dec) open CommutativeMonoid M -open EqReasoning setoid +open ≈-Reasoning setoid private variable @@ -189,7 +189,7 @@ prove′ e₁ e₂ = lemma : normalise e₁ ≡ normalise e₂ → ∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ lemma eq ρ = R.prove ρ e₁ e₂ (begin - ⟦ normalise e₁ ⟧⇓ ρ ≡⟨ P.cong (λ e → ⟦ e ⟧⇓ ρ) eq ⟩ + ⟦ normalise e₁ ⟧⇓ ρ ≡⟨ ≡.cong (λ e → ⟦ e ⟧⇓ ρ) eq ⟩ ⟦ normalise e₂ ⟧⇓ ρ ∎) -- This procedure can be combined with from-just. diff --git a/src/Algebra/Solver/IdempotentCommutativeMonoid.agda b/src/Algebra/Solver/IdempotentCommutativeMonoid.agda index fc65652787..35cfcc5ea6 100644 --- a/src/Algebra/Solver/IdempotentCommutativeMonoid.agda +++ b/src/Algebra/Solver/IdempotentCommutativeMonoid.agda @@ -20,19 +20,19 @@ open import Data.Vec.Base using (Vec; []; _∷_; lookup; replicate) open import Function.Base using (_∘_) -import Relation.Binary.Reasoning.Setoid as EqReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning import Relation.Binary.Reflection as Reflection import Relation.Nullary.Decidable as Dec import Data.Vec.Relation.Binary.Pointwise.Inductive as Pointwise -open import Relation.Binary.PropositionalEquality as P using (_≡_; decSetoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; decSetoid) open import Relation.Nullary.Decidable using (Dec) module Algebra.Solver.IdempotentCommutativeMonoid {m₁ m₂} (M : IdempotentCommutativeMonoid m₁ m₂) where open IdempotentCommutativeMonoid M -open EqReasoning setoid +open ≈-Reasoning setoid private variable @@ -202,7 +202,7 @@ prove′ e₁ e₂ = lemma : normalise e₁ ≡ normalise e₂ → ∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ lemma eq ρ = R.prove ρ e₁ e₂ (begin - ⟦ normalise e₁ ⟧⇓ ρ ≡⟨ P.cong (λ e → ⟦ e ⟧⇓ ρ) eq ⟩ + ⟦ normalise e₁ ⟧⇓ ρ ≡⟨ ≡.cong (λ e → ⟦ e ⟧⇓ ρ) eq ⟩ ⟦ normalise e₂ ⟧⇓ ρ ∎) -- This procedure can be combined with from-just. diff --git a/src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda b/src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda index a4908e851f..0ba8d9539c 100644 --- a/src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda +++ b/src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda @@ -17,8 +17,8 @@ open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (Reflexive; Sym; Trans; Antisym; Symmetric; Transitive) open import Relation.Binary.Structures using (IsEquivalence) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) -import Relation.Binary.PropositionalEquality.Properties as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +import Relation.Binary.PropositionalEquality.Properties as ≡ private variable @@ -114,16 +114,16 @@ module _ {A : Set a} where _≈_ = Pointwise _≡_ refl : Reflexive _≈_ - refl = reflexive P.refl + refl = reflexive ≡.refl sym : Symmetric _≈_ - sym = symmetric P.sym + sym = symmetric ≡.sym trans : Transitive _≈_ - trans = transitive P.trans + trans = transitive ≡.trans ≈-setoid : Setoid _ _ - ≈-setoid = setoid (P.setoid A) + ≈-setoid = setoid (≡.setoid A) ------------------------------------------------------------------------ -- Pointwise DSL @@ -161,7 +161,7 @@ module pw-Reasoning (S : Setoid a ℓ) where `head : ∀ {as bs} → `Pointwise as bs → as .head ∼ bs .head `head (`lift rs) = rs .head - `head (`refl eq) = S.reflexive (P.cong head eq) + `head (`refl eq) = S.reflexive (≡.cong head eq) `head (`bisim rs) = S.reflexive (rs .head) `head (`step `rs) = `rs .head `head (`sym `rs) = S.sym (`head `rs) @@ -169,7 +169,7 @@ module pw-Reasoning (S : Setoid a ℓ) where `tail : ∀ {as bs} → `Pointwise as bs → `Pointwise (as .tail) (bs .tail) `tail (`lift rs) = `lift (rs .tail) - `tail (`refl eq) = `refl (P.cong tail eq) + `tail (`refl eq) = `refl (≡.cong tail eq) `tail (`bisim rs) = `bisim (rs .tail) `tail (`step `rs) = `rs .tail `tail (`sym `rs) = `sym (`tail `rs) @@ -196,8 +196,8 @@ module pw-Reasoning (S : Setoid a ℓ) where pattern _≈⟨_⟨_ as bs∼as bs∼cs = `trans {as = as} (`sym (`bisim bs∼as)) bs∼cs pattern _≡⟨_⟩_ as as∼bs bs∼cs = `trans {as = as} (`refl as∼bs) bs∼cs pattern _≡⟨_⟨_ as bs∼as bs∼cs = `trans {as = as} (`sym (`refl bs∼as)) bs∼cs - pattern _≡⟨⟩_ as as∼bs = `trans {as = as} (`refl P.refl) as∼bs - pattern _∎ as = `refl {as = as} P.refl + pattern _≡⟨⟩_ as as∼bs = `trans {as = as} (`refl ≡.refl) as∼bs + pattern _∎ as = `refl {as = as} ≡.refl -- Deprecated v2.0 @@ -226,7 +226,7 @@ module pw-Reasoning (S : Setoid a ℓ) where module ≈-Reasoning {a} {A : Set a} where - open pw-Reasoning (P.setoid A) public + open pw-Reasoning (≡.setoid A) public infix 4 _≈∞_ _≈∞_ = `Pointwise∞ diff --git a/src/Codata/Musical/Colist.agda b/src/Codata/Musical/Colist.agda index 3699ef64be..c65d33e8bf 100644 --- a/src/Codata/Musical/Colist.agda +++ b/src/Codata/Musical/Colist.agda @@ -21,7 +21,7 @@ open import Data.Maybe.Relation.Unary.Any using (just) open import Data.Nat.Base using (ℕ; zero; suc) open import Data.List.Base using (List; []; _∷_) open import Data.List.NonEmpty using (List⁺; _∷_) -open import Data.Product.Base as Prod using (∃; _×_; _,_) +open import Data.Product.Base as Product using (∃; _×_; _,_) open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]′) open import Data.Vec.Bounded as Vec≤ using (Vec≤) open import Function.Base @@ -31,8 +31,8 @@ open import Relation.Binary.Core using (Rel; _⇒_) open import Relation.Binary.Bundles using (Poset; Setoid; Preorder) open import Relation.Binary.Definitions using (Transitive; Antisymmetric) import Relation.Binary.Construct.FromRel as Ind -import Relation.Binary.Reasoning.Preorder as PreR -import Relation.Binary.Reasoning.PartialOrder as POR +import Relation.Binary.Reasoning.Preorder as ≲-Reasoning +import Relation.Binary.Reasoning.PartialOrder as ≤-Reasoning open import Relation.Binary.PropositionalEquality as P using (_≡_) open import Relation.Binary.Reasoning.Syntax open import Relation.Nullary.Reflects using (invert) @@ -108,7 +108,7 @@ Any-∈ {P = P} = mk↔ₛ′ 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) + to (there p) = Product.map id (Product.map there id) (to p) from : ∀ {x xs} → x ∈ xs → P x → Any P xs from (here P.refl) p = here p @@ -118,7 +118,7 @@ Any-∈ {P = P} = mk↔ₛ′ 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) + P.cong (Product.map id (Product.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 @@ -164,7 +164,7 @@ Any-∈ {P = P} = mk↔ₛ′ antisym (x ∷ p₁) p₂ = x ∷ ♯ antisym (♭ p₁) (tail p₂) module ⊑-Reasoning {a} {A : Set a} where - private module Base = POR (⊑-Poset A) + private module Base = ≤-Reasoning (⊑-Poset A) open Base public hiding (step-<; step-≤) @@ -188,7 +188,7 @@ module ⊑-Reasoning {a} {A : Set a} where -- ys ≡⟨ ys≡zs ⟩ -- zs ∎ module ⊆-Reasoning {A : Set a} where - private module Base = PreR (⊆-Preorder A) + private module Base = ≲-Reasoning (⊆-Preorder A) open Base public hiding (step-≲; step-∼) diff --git a/src/Codata/Musical/Colist/Infinite-merge.agda b/src/Codata/Musical/Colist/Infinite-merge.agda index 584880a768..ecf393ca83 100644 --- a/src/Codata/Musical/Colist/Infinite-merge.agda +++ b/src/Codata/Musical/Colist/Infinite-merge.agda @@ -13,7 +13,7 @@ open import Codata.Musical.Colist as Colist hiding (_⋎_) open import Data.Nat.Base open import Data.Nat.Induction using (<′-wellFounded) open import Data.Nat.Properties -open import Data.Product.Base as Prod using (_×_; _,_; ∃; ∃₂; proj₁; proj₂) +open import Data.Product.Base as Product using (_×_; _,_; ∃; ∃₂; proj₁; proj₂) open import Data.Sum.Base open import Data.Sum.Properties open import Data.Sum.Function.Propositional using (_⊎-cong_) @@ -195,7 +195,7 @@ Any-merge {P = P} xss = mk↔ₛ′ (proj₁ ∘ to xss) from to∘from (proj₂ | index-Any-⋎P xs p ... | inj₁ q | P.refl | _ = here (inj₂ q) , P.refl ... | inj₂ q | P.refl | q≤p = - Prod.map there + Product.map there (P.cong (there ∘ (Inverse.from (Any-⋎P xs)) ∘ inj₂)) (rec (s≤′s q≤p)) diff --git a/src/Codata/Sized/Colist/Properties.agda b/src/Codata/Sized/Colist/Properties.agda index a4196bb952..c03b3c2096 100644 --- a/src/Codata/Sized/Colist/Properties.agda +++ b/src/Codata/Sized/Colist/Properties.agda @@ -14,24 +14,24 @@ open import Codata.Sized.Thunk as Thunk using (Thunk; force) open import Codata.Sized.Colist open import Codata.Sized.Colist.Bisimilarity open import Codata.Sized.Conat -open import Codata.Sized.Conat.Bisimilarity as coℕᵇ using (zero; suc) -import Codata.Sized.Conat.Properties as coℕₚ +open import Codata.Sized.Conat.Bisimilarity as Conat using (zero; suc) +import Codata.Sized.Conat.Properties as Conat open import Codata.Sized.Cowriter as Cowriter using ([_]; _∷_) -open import Codata.Sized.Cowriter.Bisimilarity as coWriterᵇ using ([_]; _∷_) +open import Codata.Sized.Cowriter.Bisimilarity as Cowriter using ([_]; _∷_) open import Codata.Sized.Stream as Stream using (Stream; _∷_) open import Data.Vec.Bounded as Vec≤ using (Vec≤) open import Data.List.Base as List using (List; []; _∷_) open import Data.List.NonEmpty as List⁺ using (List⁺; _∷_) open import Data.List.Relation.Binary.Equality.Propositional using (≋-refl) open import Data.Maybe.Base as Maybe using (Maybe; nothing; just) -import Data.Maybe.Properties as Maybeₚ +import Data.Maybe.Properties as Maybe open import Data.Maybe.Relation.Unary.All using (All; nothing; just) open import Data.Nat.Base as ℕ using (zero; suc; z≤n; s≤s) -open import Data.Product.Base as Prod using (_×_; _,_; uncurry) +open import Data.Product.Base as Product using (_×_; _,_; uncurry) open import Data.These.Base as These using (These; this; that; these) open import Data.Vec.Base as Vec using (Vec; []; _∷_) open import Function.Base -open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) private variable @@ -47,13 +47,13 @@ private map-id : ∀ (as : Colist A ∞) → i ⊢ map id as ≈ as map-id [] = [] -map-id (a ∷ as) = Eq.refl ∷ λ where .force → map-id (as .force) +map-id (a ∷ as) = ≡.refl ∷ λ where .force → map-id (as .force) map-∘ : ∀ (f : A → B) (g : B → C) as {i} → i ⊢ map g (map f as) ≈ map (g ∘ f) as map-∘ f g [] = [] map-∘ f g (a ∷ as) = - Eq.refl ∷ λ where .force → map-∘ f g (as .force) + ≡.refl ∷ λ where .force → map-∘ f g (as .force) ------------------------------------------------------------------------ -- Relation to Cowriter @@ -62,24 +62,24 @@ fromCowriter∘toCowriter≗id : ∀ (as : Colist A ∞) → i ⊢ fromCowriter (toCowriter as) ≈ as fromCowriter∘toCowriter≗id [] = [] fromCowriter∘toCowriter≗id (a ∷ as) = - Eq.refl ∷ λ where .force → fromCowriter∘toCowriter≗id (as .force) + ≡.refl ∷ λ where .force → fromCowriter∘toCowriter≗id (as .force) ------------------------------------------------------------------------ -- Properties of length -length-∷ : ∀ (a : A) as → i coℕᵇ.⊢ length (a ∷ as) ≈ 1 ℕ+ length (as .force) -length-∷ a as = suc (λ where .force → coℕᵇ.refl) +length-∷ : ∀ (a : A) as → i Conat.⊢ length (a ∷ as) ≈ 1 ℕ+ length (as .force) +length-∷ a as = suc (λ where .force → Conat.refl) -length-replicate : ∀ n (a : A) → i coℕᵇ.⊢ length (replicate n a) ≈ n +length-replicate : ∀ n (a : A) → i Conat.⊢ length (replicate n a) ≈ n length-replicate zero a = zero length-replicate (suc n) a = suc λ where .force → length-replicate (n .force) a length-++ : (as bs : Colist A ∞) → - i coℕᵇ.⊢ length (as ++ bs) ≈ length as + length bs -length-++ [] bs = coℕᵇ.refl + i Conat.⊢ length (as ++ bs) ≈ length as + length bs +length-++ [] bs = Conat.refl length-++ (a ∷ as) bs = suc λ where .force → length-++ (as .force) bs -length-map : ∀ (f : A → B) as → i coℕᵇ.⊢ length (map f as) ≈ length as +length-map : ∀ (f : A → B) as → i Conat.⊢ length (map f as) ≈ length as length-map f [] = zero length-map f (a ∷ as) = suc λ where .force → length-map f (as .force) @@ -89,51 +89,51 @@ length-map f (a ∷ as) = suc λ where .force → length-map f (as .force) replicate-+ : ∀ m n (a : A) → i ⊢ replicate (m + n) a ≈ replicate m a ++ replicate n a replicate-+ zero n a = refl -replicate-+ (suc m) n a = Eq.refl ∷ λ where .force → replicate-+ (m .force) n a +replicate-+ (suc m) n a = ≡.refl ∷ λ where .force → replicate-+ (m .force) n a map-replicate : ∀ (f : A → B) n a → i ⊢ map f (replicate n a) ≈ replicate n (f a) map-replicate f zero a = [] map-replicate f (suc n) a = - Eq.refl ∷ λ where .force → map-replicate f (n .force) a + ≡.refl ∷ λ where .force → map-replicate f (n .force) a lookup-replicate : ∀ k n (a : A) → All (a ≡_) (lookup (replicate n a) k) lookup-replicate k zero a = nothing -lookup-replicate zero (suc n) a = just Eq.refl +lookup-replicate zero (suc n) a = just ≡.refl lookup-replicate (suc k) (suc n) a = lookup-replicate k (n .force) a ------------------------------------------------------------------------ -- Properties of unfold map-unfold : ∀ (f : B → C) (alg : A → Maybe (A × B)) a → - i ⊢ map f (unfold alg a) ≈ unfold (Maybe.map (Prod.map₂ f) ∘ alg) a + i ⊢ map f (unfold alg a) ≈ unfold (Maybe.map (Product.map₂ f) ∘ alg) a map-unfold f alg a with alg a ... | nothing = [] -... | just (a′ , b) = Eq.refl ∷ λ where .force → map-unfold f alg a′ +... | just (a′ , b) = ≡.refl ∷ λ where .force → map-unfold f alg a′ module _ {alg : A → Maybe (A × B)} {a} where unfold-nothing : alg a ≡ nothing → unfold alg a ≡ [] unfold-nothing eq with alg a - ... | nothing = Eq.refl + ... | nothing = ≡.refl unfold-just : ∀ {a′ b} → alg a ≡ just (a′ , b) → i ⊢ unfold alg a ≈ b ∷ λ where .force → unfold alg a′ unfold-just eq with alg a - unfold-just Eq.refl | just (a′ , b) = Eq.refl ∷ λ where .force → refl + unfold-just ≡.refl | just (a′ , b) = ≡.refl ∷ λ where .force → refl ------------------------------------------------------------------------ -- Properties of scanl length-scanl : ∀ (c : B → A → B) n as → - i coℕᵇ.⊢ length (scanl c n as) ≈ 1 ℕ+ length as + i Conat.⊢ length (scanl c n as) ≈ 1 ℕ+ length as length-scanl c n [] = suc λ where .force → zero length-scanl c n (a ∷ as) = suc λ { .force → begin length (scanl c (c n a) (as .force)) ≈⟨ length-scanl c (c n a) (as .force) ⟩ 1 ℕ+ length (as .force) ≈⟨ length-∷ a as ⟨ - length (a ∷ as) ∎ } where open coℕᵇ.≈-Reasoning + length (a ∷ as) ∎ } where open Conat.≈-Reasoning module _ (cons : C → B → C) (alg : A → Maybe (A × B)) where @@ -145,15 +145,15 @@ module _ (cons : C → B → C) (alg : A → Maybe (A × B)) where scanl-unfold : ∀ nil a → i ⊢ scanl cons nil (unfold alg a) ≈ nil ∷ (λ where .force → unfold alg′ (a , nil)) scanl-unfold nil a with alg a in eq - ... | nothing = Eq.refl ∷ λ { .force → - sym (fromEq (unfold-nothing (Maybeₚ.map-nothing eq))) } - ... | just (a′ , b) = Eq.refl ∷ λ { .force → begin + ... | nothing = ≡.refl ∷ λ { .force → + sym (fromEq (unfold-nothing (Maybe.map-nothing eq))) } + ... | just (a′ , b) = ≡.refl ∷ λ { .force → begin scanl cons (cons nil b) (unfold alg a′) ≈⟨ scanl-unfold (cons nil b) a′ ⟩ (cons nil b ∷ _) - ≈⟨ Eq.refl ∷ (λ where .force → refl) ⟩ + ≈⟨ ≡.refl ∷ (λ where .force → refl) ⟩ (cons nil b ∷ _) - ≈⟨ unfold-just (Maybeₚ.map-just eq) ⟨ + ≈⟨ unfold-just (Maybe.map-just eq) ⟨ unfold alg′ (a , nil) ∎ } where open ≈-Reasoning ------------------------------------------------------------------------ @@ -164,10 +164,10 @@ map-alignWith : ∀ (f : C → D) (al : These A B → C) as bs → map-alignWith f al [] bs = map-∘ (al ∘′ that) f bs map-alignWith f al as@(_ ∷ _) [] = map-∘ (al ∘′ this) f as map-alignWith f al (a ∷ as) (b ∷ bs) = - Eq.refl ∷ λ where .force → map-alignWith f al (as .force) (bs .force) + ≡.refl ∷ λ where .force → map-alignWith f al (as .force) (bs .force) length-alignWith : ∀ (al : These A B → C) as bs → - i coℕᵇ.⊢ length (alignWith al as bs) ≈ length as ⊔ length bs + i Conat.⊢ length (alignWith al as bs) ≈ length as ⊔ length bs length-alignWith al [] bs = length-map (al ∘ that) bs length-alignWith al as@(_ ∷ _) [] = length-map (al ∘ this) as length-alignWith al (a ∷ as) (b ∷ bs) = @@ -181,10 +181,10 @@ map-zipWith : ∀ (f : C → D) (zp : A → B → C) as bs → map-zipWith f zp [] _ = [] map-zipWith f zp (_ ∷ _) [] = [] map-zipWith f zp (a ∷ as) (b ∷ bs) = - Eq.refl ∷ λ where .force → map-zipWith f zp (as .force) (bs .force) + ≡.refl ∷ λ where .force → map-zipWith f zp (as .force) (bs .force) length-zipWith : ∀ (zp : A → B → C) as bs → - i coℕᵇ.⊢ length (zipWith zp as bs) ≈ length as ⊓ length bs + i Conat.⊢ length (zipWith zp as bs) ≈ length as ⊓ length bs length-zipWith zp [] bs = zero length-zipWith zp as@(_ ∷ _) [] = zero length-zipWith zp (a ∷ as) (b ∷ bs) = @@ -208,25 +208,25 @@ map-drop f zero as = refl map-drop f (suc m) [] = [] map-drop f (suc m) (a ∷ as) = map-drop f m (as .force) -length-drop : ∀ m (as : Colist A ∞) → i coℕᵇ.⊢ length (drop m as) ≈ length as ∸ m -length-drop zero as = coℕᵇ.refl -length-drop (suc m) [] = coℕᵇ.sym (coℕₚ.0∸m≈0 m) +length-drop : ∀ m (as : Colist A ∞) → i Conat.⊢ length (drop m as) ≈ length as ∸ m +length-drop zero as = Conat.refl +length-drop (suc m) [] = Conat.sym (Conat.0∸m≈0 m) length-drop (suc m) (a ∷ as) = length-drop m (as .force) drop-fromList-++-identity : ∀ (as : List A) bs → drop (List.length as) (fromList as ++ bs) ≡ bs -drop-fromList-++-identity [] bs = Eq.refl +drop-fromList-++-identity [] bs = ≡.refl drop-fromList-++-identity (a ∷ as) bs = drop-fromList-++-identity as bs drop-fromList-++-≤ : ∀ (as : List A) bs {m} → m ℕ.≤ List.length as → drop m (fromList as ++ bs) ≡ fromList (List.drop m as) ++ bs -drop-fromList-++-≤ [] bs z≤n = Eq.refl -drop-fromList-++-≤ (a ∷ as) bs z≤n = Eq.refl +drop-fromList-++-≤ [] bs z≤n = ≡.refl +drop-fromList-++-≤ (a ∷ as) bs z≤n = ≡.refl drop-fromList-++-≤ (a ∷ as) bs (s≤s p) = drop-fromList-++-≤ as bs p drop-fromList-++-≥ : ∀ (as : List A) bs {m} → m ℕ.≥ List.length as → drop m (fromList as ++ bs) ≡ drop (m ℕ.∸ List.length as) bs -drop-fromList-++-≥ [] bs z≤n = Eq.refl +drop-fromList-++-≥ [] bs z≤n = ≡.refl drop-fromList-++-≥ (a ∷ as) bs (s≤s p) = drop-fromList-++-≥ as bs p drop-⁺++-identity : ∀ (as : List⁺ A) bs → @@ -236,7 +236,7 @@ drop-⁺++-identity (a ∷ as) bs = drop-fromList-++-identity as (bs .force) ------------------------------------------------------------------------ -- Properties of cotake -length-cotake : ∀ n (as : Stream A ∞) → i coℕᵇ.⊢ length (cotake n as) ≈ n +length-cotake : ∀ n (as : Stream A ∞) → i Conat.⊢ length (cotake n as) ≈ n length-cotake zero as = zero length-cotake (suc n) (a ∷ as) = suc λ where .force → length-cotake (n .force) (as .force) @@ -245,7 +245,7 @@ map-cotake : ∀ (f : A → B) n as → i ⊢ map f (cotake n as) ≈ cotake n (Stream.map f as) map-cotake f zero as = [] map-cotake f (suc n) (a ∷ as) = - Eq.refl ∷ λ where .force → map-cotake f (n .force) (as .force) + ≡.refl ∷ λ where .force → map-cotake f (n .force) (as .force) ------------------------------------------------------------------------ -- Properties of chunksOf @@ -255,20 +255,20 @@ module Map-ChunksOf (f : A → B) n where open ChunksOf n using (chunksOfAcc) map-chunksOf : ∀ as → - i coWriterᵇ.⊢ Cowriter.map (Vec.map f) (Vec≤.map f) (chunksOf n as) + i Cowriter.⊢ Cowriter.map (Vec.map f) (Vec≤.map f) (chunksOf n as) ≈ chunksOf n (map f as) map-chunksOfAcc : ∀ m as {k≤ k≡ k≤′ k≡′} → (∀ vs → Vec≤.map f (k≤ vs) ≡ k≤′ (Vec≤.map f vs)) → (∀ vs → Vec.map f (k≡ vs) ≡ k≡′ (Vec.map f vs)) → - i coWriterᵇ.⊢ Cowriter.map (Vec.map f) (Vec≤.map f) + i Cowriter.⊢ Cowriter.map (Vec.map f) (Vec≤.map f) (chunksOfAcc m k≤ k≡ as) ≈ chunksOfAcc m k≤′ k≡′ (map f as) - map-chunksOf as = map-chunksOfAcc n as (λ vs → Eq.refl) (λ vs → Eq.refl) + map-chunksOf as = map-chunksOfAcc n as (λ vs → ≡.refl) (λ vs → ≡.refl) map-chunksOfAcc zero as eq-≤ eq-≡ = eq-≡ [] ∷ λ where .force → map-chunksOf as - map-chunksOfAcc (suc m) [] eq-≤ eq-≡ = coWriterᵇ.[ eq-≤ Vec≤.[] ] + map-chunksOfAcc (suc m) [] eq-≤ eq-≡ = Cowriter.[ eq-≤ Vec≤.[] ] map-chunksOfAcc (suc m) (a ∷ as) eq-≤ eq-≡ = map-chunksOfAcc m (as .force) (eq-≤ ∘ (a Vec≤.∷_)) (eq-≡ ∘ (a Vec.∷_)) @@ -280,21 +280,21 @@ open Map-ChunksOf using (map-chunksOf) public fromList-++ : (as bs : List A) → i ⊢ fromList (as List.++ bs) ≈ fromList as ++ fromList bs fromList-++ [] bs = refl -fromList-++ (a ∷ as) bs = Eq.refl ∷ λ where .force → fromList-++ as bs +fromList-++ (a ∷ as) bs = ≡.refl ∷ λ where .force → fromList-++ as bs fromList-scanl : ∀ (c : B → A → B) n as → i ⊢ fromList (List.scanl c n as) ≈ scanl c n (fromList as) -fromList-scanl c n [] = Eq.refl ∷ λ where .force → refl +fromList-scanl c n [] = ≡.refl ∷ λ where .force → refl fromList-scanl c n (a ∷ as) = - Eq.refl ∷ λ where .force → fromList-scanl c (c n a) as + ≡.refl ∷ λ where .force → fromList-scanl c (c n a) as map-fromList : ∀ (f : A → B) as → i ⊢ map f (fromList as) ≈ fromList (List.map f as) map-fromList f [] = [] -map-fromList f (a ∷ as) = Eq.refl ∷ λ where .force → map-fromList f as +map-fromList f (a ∷ as) = ≡.refl ∷ λ where .force → map-fromList f as length-fromList : (as : List A) → - i coℕᵇ.⊢ length (fromList as) ≈ fromℕ (List.length as) + i Conat.⊢ length (fromList as) ≈ fromℕ (List.length as) length-fromList [] = zero length-fromList (a ∷ as) = suc (λ where .force → length-fromList as) @@ -304,19 +304,19 @@ length-fromList (a ∷ as) = suc (λ where .force → length-fromList as) fromStream-++ : ∀ (as : List A) bs → i ⊢ fromStream (as Stream.++ bs) ≈ fromList as ++ fromStream bs fromStream-++ [] bs = refl -fromStream-++ (a ∷ as) bs = Eq.refl ∷ λ where .force → fromStream-++ as bs +fromStream-++ (a ∷ as) bs = ≡.refl ∷ λ where .force → fromStream-++ as bs fromStream-⁺++ : ∀ (as : List⁺ A) bs → i ⊢ fromStream (as Stream.⁺++ bs) ≈ fromList⁺ as ++ fromStream (bs .force) fromStream-⁺++ (a ∷ as) bs = - Eq.refl ∷ λ where .force → fromStream-++ as (bs .force) + ≡.refl ∷ λ where .force → fromStream-++ as (bs .force) fromStream-concat : (ass : Stream (List⁺ A) ∞) → i ⊢ concat (fromStream ass) ≈ fromStream (Stream.concat ass) fromStream-concat (as@(a ∷ _) ∷ ass) = begin concat (fromStream (as ∷ ass)) - ≈⟨ Eq.refl ∷ (λ { .force → ++⁺ ≋-refl (fromStream-concat (ass .force))}) ⟩ + ≈⟨ ≡.refl ∷ (λ { .force → ++⁺ ≋-refl (fromStream-concat (ass .force))}) ⟩ a ∷ _ ≈⟨ sym (fromStream-⁺++ as _) ⟩ fromStream (Stream.concat (as ∷ ass)) ∎ where open ≈-Reasoning @@ -325,12 +325,12 @@ fromStream-scanl : ∀ (c : B → A → B) n as → i ⊢ scanl c n (fromStream as) ≈ fromStream (Stream.scanl c n as) fromStream-scanl c n (a ∷ as) = - Eq.refl ∷ λ where .force → fromStream-scanl c (c n a) (as .force) + ≡.refl ∷ λ where .force → fromStream-scanl c (c n a) (as .force) map-fromStream : ∀ (f : A → B) as → i ⊢ map f (fromStream as) ≈ fromStream (Stream.map f as) map-fromStream f (a ∷ as) = - Eq.refl ∷ λ where .force → map-fromStream f (as .force) + ≡.refl ∷ λ where .force → map-fromStream f (as .force) ------------------------------------------------------------------------ -- DEPRECATED diff --git a/src/Codata/Sized/Cowriter.agda b/src/Codata/Sized/Cowriter.agda index 7206c83c86..03d24d3d3c 100644 --- a/src/Codata/Sized/Cowriter.agda +++ b/src/Codata/Sized/Cowriter.agda @@ -17,8 +17,8 @@ open import Codata.Sized.Stream as Stream using (Stream; _∷_) open import Data.Unit.Base open import Data.List.Base using (List; []; _∷_) open import Data.List.NonEmpty.Base using (List⁺; _∷_) -open import Data.Nat.Base as Nat using (ℕ; zero; suc) -open import Data.Product.Base as Prod using (_×_; _,_) +open import Data.Nat.Base as ℕ using (ℕ; zero; suc) +open import Data.Product.Base as Product using (_×_; _,_) open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂) open import Data.Vec.Base using (Vec; []; _∷_) open import Data.Vec.Bounded.Base as Vec≤ using (Vec≤; _,_) @@ -68,11 +68,11 @@ length (w ∷ cw) = suc λ where .force → length (cw .force) splitAt : ∀ (n : ℕ) → Cowriter W A ∞ → (Vec W n × Cowriter W A ∞) ⊎ (Vec≤ W n × A) splitAt zero cw = inj₁ ([] , cw) splitAt (suc n) [ a ] = inj₂ (Vec≤.[] , a) -splitAt (suc n) (w ∷ cw) = Sum.map (Prod.map₁ (w ∷_)) (Prod.map₁ (w Vec≤.∷_)) +splitAt (suc n) (w ∷ cw) = Sum.map (Product.map₁ (w ∷_)) (Product.map₁ (w Vec≤.∷_)) $ splitAt n (cw .force) take : ∀ (n : ℕ) → Cowriter W A ∞ → Vec W n ⊎ (Vec≤ W n × A) -take n = Sum.map₁ Prod.proj₁ ∘′ splitAt n +take n = Sum.map₁ Product.proj₁ ∘′ splitAt n infixr 5 _++_ _⁺++_ _++_ : ∀ {i} → List W → Cowriter W A i → Cowriter W A i diff --git a/src/Codata/Sized/M/Properties.agda b/src/Codata/Sized/M/Properties.agda index f9801448f5..a02410d43e 100644 --- a/src/Codata/Sized/M/Properties.agda +++ b/src/Codata/Sized/M/Properties.agda @@ -15,7 +15,7 @@ open import Codata.Sized.M open import Codata.Sized.M.Bisimilarity open import Data.Container.Core as C hiding (map) import Data.Container.Morphism as Mp -open import Data.Product.Base as Prod using (_,_) +open import Data.Product.Base as Product using (_,_) open import Data.Product.Properties hiding (map-cong) open import Function.Base using (_$′_; _∘′_) import Relation.Binary.PropositionalEquality.Core as P diff --git a/src/Codata/Sized/Stream/Properties.agda b/src/Codata/Sized/Stream/Properties.agda index 40cb3fc961..0fd436f1bd 100644 --- a/src/Codata/Sized/Stream/Properties.agda +++ b/src/Codata/Sized/Stream/Properties.agda @@ -20,7 +20,7 @@ open import Data.Nat.GeneralisedArithmetic using (fold; fold-pull) open import Data.List.Base as List using ([]; _∷_) open import Data.List.NonEmpty as List⁺ using (List⁺; _∷_) import Data.List.Relation.Binary.Equality.Propositional as Eq -open import Data.Product.Base as Prod using (_,_) +open import Data.Product.Base as Product using (_,_) open import Data.Vec.Base as Vec using (_∷_) open import Function.Base using (id; _$_; _∘′_; const) @@ -49,7 +49,7 @@ take-repeat-identity (suc n) a = P.cong (a Vec.∷_) (take-repeat-identity n a) splitAt-repeat-identity : (n : ℕ) (a : A) → splitAt n (repeat a) ≡ (Vec.replicate n a , repeat a) splitAt-repeat-identity zero a = P.refl -splitAt-repeat-identity (suc n) a = P.cong (Prod.map₁ (a ∷_)) (splitAt-repeat-identity n a) +splitAt-repeat-identity (suc n) a = P.cong (Product.map₁ (a ∷_)) (splitAt-repeat-identity n a) replicate-repeat : ∀ {i} (n : ℕ) (a : A) → i ⊢ List.replicate n a ++ repeat a ≈ repeat a replicate-repeat zero a = refl @@ -103,10 +103,10 @@ map-∘ f g (a ∷ as) = P.refl ∷ λ where .force → map-∘ f g (as .force) -- splitAt splitAt-map : ∀ n (f : A → B) xs → - splitAt n (map f xs) ≡ Prod.map (Vec.map f) (map f) (splitAt n xs) + splitAt n (map f xs) ≡ Product.map (Vec.map f) (map f) (splitAt n xs) splitAt-map zero f xs = P.refl splitAt-map (suc n) f (x ∷ xs) = - P.cong (Prod.map₁ (f x Vec.∷_)) (splitAt-map n f (xs .force)) + P.cong (Product.map₁ (f x Vec.∷_)) (splitAt-map n f (xs .force)) ------------------------------------------------------------------------ -- iterate diff --git a/src/Data/Char/Properties.agda b/src/Data/Char/Properties.agda index c399bf30e3..b9c7270881 100644 --- a/src/Data/Char/Properties.agda +++ b/src/Data/Char/Properties.agda @@ -11,7 +11,7 @@ module Data.Char.Properties where open import Data.Bool.Base using (Bool) open import Data.Char.Base import Data.Nat.Base as ℕ -import Data.Nat.Properties as ℕₚ +import Data.Nat.Properties as ℕ open import Data.Product.Base using (_,_) open import Function.Base @@ -27,10 +27,10 @@ open import Relation.Binary.Definitions import Relation.Binary.Construct.On as On import Relation.Binary.Construct.Subst.Equality as Subst import Relation.Binary.Construct.Closure.Reflexive as Refl -import Relation.Binary.Construct.Closure.Reflexive.Properties as Reflₚ -open import Relation.Binary.PropositionalEquality.Core as PropEq +import Relation.Binary.Construct.Closure.Reflexive.Properties as Refl +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; _≢_; refl; cong; sym; trans; subst) -import Relation.Binary.PropositionalEquality.Properties as PropEq +import Relation.Binary.PropositionalEquality.Properties as ≡ ------------------------------------------------------------------------ -- Primitive properties @@ -56,16 +56,16 @@ open import Agda.Builtin.Char.Properties infix 4 _≟_ _≟_ : Decidable {A = Char} _≡_ -x ≟ y = map′ ≈⇒≡ ≈-reflexive (toℕ x ℕₚ.≟ toℕ y) +x ≟ y = map′ ≈⇒≡ ≈-reflexive (toℕ x ℕ.≟ toℕ y) setoid : Setoid _ _ -setoid = PropEq.setoid Char +setoid = ≡.setoid Char decSetoid : DecSetoid _ _ -decSetoid = PropEq.decSetoid _≟_ +decSetoid = ≡.decSetoid _≟_ isDecEquivalence : IsDecEquivalence _≡_ -isDecEquivalence = PropEq.isDecEquivalence _≟_ +isDecEquivalence = ≡.isDecEquivalence _≟_ ------------------------------------------------------------------------ -- Boolean equality test. @@ -95,30 +95,30 @@ private infix 4 _ ¬lt ¬eq gt = tri> ¬lt (≉⇒≢ ¬eq) gt <-irrefl : Irreflexive _≡_ _<_ -<-irrefl = ℕₚ.<-irrefl ∘′ cong toℕ +<-irrefl = ℕ.<-irrefl ∘′ cong toℕ <-trans : Transitive _<_ -<-trans {c} {d} {e} = On.transitive toℕ ℕ._<_ ℕₚ.<-trans {c} {d} {e} +<-trans {c} {d} {e} = On.transitive toℕ ℕ._<_ ℕ.<-trans {c} {d} {e} <-asym : Asymmetric _<_ -<-asym {c} {d} = On.asymmetric toℕ ℕ._<_ ℕₚ.<-asym {c} {d} +<-asym {c} {d} = On.asymmetric toℕ ℕ._<_ ℕ.<-asym {c} {d} <-isStrictPartialOrder : IsStrictPartialOrder _≡_ _<_ <-isStrictPartialOrder = record - { isEquivalence = PropEq.isEquivalence + { isEquivalence = ≡.isEquivalence ; irrefl = <-irrefl ; trans = λ {a} {b} {c} → <-trans {a} {b} {c} - ; <-resp-≈ = (λ {c} → PropEq.subst (c <_)) - , (λ {c} → PropEq.subst (_< c)) + ; <-resp-≈ = (λ {c} → ≡.subst (c <_)) + , (λ {c} → ≡.subst (_< c)) } <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ @@ -142,20 +142,20 @@ _) open import Data.Product.Properties using (,-injective) open import Data.Product.Algebra using (×-cong) @@ -174,14 +174,14 @@ toℕ≤pred[n] zero = z≤n toℕ≤pred[n] (suc {n = suc n} i) = s≤s (toℕ≤pred[n] i) toℕ≤n : ∀ (i : Fin n) → toℕ i ℕ.≤ n -toℕ≤n {suc n} i = ℕₚ.m≤n⇒m≤1+n (toℕ≤pred[n] i) +toℕ≤n {suc n} i = ℕ.m≤n⇒m≤1+n (toℕ≤pred[n] i) -- A simpler implementation of toℕ≤pred[n], -- however, with a different reduction behavior. -- If no one needs the reduction behavior of toℕ≤pred[n], -- it can be removed in favor of toℕ≤pred[n]′. toℕ≤pred[n]′ : ∀ (i : Fin n) → toℕ i ℕ.≤ ℕ.pred n -toℕ≤pred[n]′ i = ℕₚ.<⇒≤pred (toℕ>=_ @@ -307,13 +307,13 @@ module _ (_≈?_ : B.Decidable {A = A} _≡_) where -- length ∈-length : ∀ {x : A} {xs} → x ∈ xs → 1 ≤ length xs -∈-length = Membershipₛ.∈-length (P.setoid _) +∈-length = Membership.∈-length (≡.setoid _) ------------------------------------------------------------------------ -- lookup ∈-lookup : ∀ {xs : List A} i → lookup xs i ∈ xs -∈-lookup {xs = xs} i = Membershipₛ.∈-lookup (P.setoid _) xs i +∈-lookup {xs = xs} i = Membership.∈-lookup (≡.setoid _) xs i ------------------------------------------------------------------------ -- foldr @@ -322,7 +322,7 @@ module _ {_•_ : Op₂ A} where foldr-selective : Selective _≡_ _•_ → ∀ e xs → (foldr _•_ e xs ≡ e) ⊎ (foldr _•_ e xs ∈ xs) - foldr-selective = Membershipₛ.foldr-selective (P.setoid A) + foldr-selective = Membership.foldr-selective (≡.setoid A) ------------------------------------------------------------------------ -- allFin @@ -380,7 +380,7 @@ finite inj (x ∷ xs) fᵢ∈x∷xs = ¬¬-excluded-middle helper f′-injective′ : Injective _≡_ _≡_ f′ f′-injective′ {j} {k} eq with i ≤ᵇ j | Reflects.invert (≤ᵇ-reflects-≤ i j) | i ≤ᵇ k | Reflects.invert (≤ᵇ-reflects-≤ i k) - ... | true | p | true | q = P.cong pred (f-inj eq) + ... | true | p | true | q = ≡.cong pred (f-inj eq) ... | true | p | false | q = contradiction (f-inj eq) (lemma p q) ... | false | p | true | q = contradiction (f-inj eq) (lemma q p ∘ sym) ... | false | p | false | q = f-inj eq @@ -388,7 +388,7 @@ finite inj (x ∷ xs) fᵢ∈x∷xs = ¬¬-excluded-middle helper f′-inj : ℕ ↣ _ f′-inj = record { to = f′ - ; cong = P.cong f′ + ; cong = ≡.cong f′ ; injective = f′-injective′ } @@ -398,4 +398,4 @@ finite inj (x ∷ xs) fᵢ∈x∷xs = ¬¬-excluded-middle helper there-injective-≢∈ : ∀ {xs} {x y z : A} {x∈xs : x ∈ xs} {y∈xs : y ∈ xs} → there {x = z} x∈xs ≢∈ there y∈xs → x∈xs ≢∈ y∈xs -there-injective-≢∈ neq refl eq = neq refl (P.cong there eq) +there-injective-≢∈ neq refl eq = neq refl (≡.cong there eq) diff --git a/src/Data/List/Membership/Propositional/Properties/Core.agda b/src/Data/List/Membership/Propositional/Properties/Core.agda index fadd49bb4b..f7543d05f8 100644 --- a/src/Data/List/Membership/Propositional/Properties/Core.agda +++ b/src/Data/List/Membership/Propositional/Properties/Core.agda @@ -17,7 +17,7 @@ open import Function.Bundles open import Data.List.Base using (List) open import Data.List.Relation.Unary.Any as Any using (Any; here; there) open import Data.List.Membership.Propositional -open import Data.Product.Base as Prod +open import Data.Product.Base as Product using (_,_; proj₁; proj₂; uncurry′; ∃; _×_) open import Level using (Level) open import Relation.Binary.PropositionalEquality.Core as P @@ -42,7 +42,7 @@ map∘find (there p) hyp = P.cong there (map∘find p hyp) find∘map : ∀ {P : Pred A p} {Q : Pred A 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 (Any.map f p) ≡ Product.map id (Product.map id f) (find p) find∘map (here p) f = refl find∘map (there p) f rewrite find∘map p f = refl diff --git a/src/Data/List/Membership/Propositional/Properties/WithK.agda b/src/Data/List/Membership/Propositional/Properties/WithK.agda index 60a86ee5e3..a86f23a8ea 100644 --- a/src/Data/List/Membership/Propositional/Properties/WithK.agda +++ b/src/Data/List/Membership/Propositional/Properties/WithK.agda @@ -12,9 +12,9 @@ module Data.List.Membership.Propositional.Properties.WithK where open import Data.List.Base open import Data.List.Relation.Unary.Unique.Propositional open import Data.List.Membership.Propositional -import Data.List.Membership.Setoid.Properties as Membershipₛ +import Data.List.Membership.Setoid.Properties as Membership open import Relation.Unary using (Irrelevant) -open import Relation.Binary.PropositionalEquality.Properties as P +open import Relation.Binary.PropositionalEquality.Properties as ≡ open import Relation.Binary.PropositionalEquality.WithK ------------------------------------------------------------------------ @@ -22,4 +22,4 @@ open import Relation.Binary.PropositionalEquality.WithK unique⇒irrelevant : ∀ {a} {A : Set a} {xs : List A} → Unique xs → Irrelevant (_∈ xs) -unique⇒irrelevant = Membershipₛ.unique⇒irrelevant (P.setoid _) ≡-irrelevant +unique⇒irrelevant = Membership.unique⇒irrelevant (≡.setoid _) ≡-irrelevant diff --git a/src/Data/List/Membership/Setoid.agda b/src/Data/List/Membership/Setoid.agda index 1a19169992..cadc600646 100644 --- a/src/Data/List/Membership/Setoid.agda +++ b/src/Data/List/Membership/Setoid.agda @@ -15,7 +15,7 @@ open import Function.Base using (_∘_; id; flip) open import Data.List.Base as List using (List; []; _∷_; length; lookup) open import Data.List.Relation.Unary.Any as Any using (Any; index; map; here; there) -open import Data.Product.Base as Prod using (∃; _×_; _,_) +open import Data.Product.Base as Product using (∃; _×_; _,_) open import Relation.Unary using (Pred) open import Relation.Nullary.Negation using (¬_) @@ -50,7 +50,7 @@ module _ {p} {P : Pred A p} where find : ∀ {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) + find (there pxs) = Product.map id (Product.map there id) (find pxs) lose : P Respects _≈_ → ∀ {x xs} → x ∈ xs → P x → Any P xs lose resp x∈xs px = map (flip resp px) x∈xs diff --git a/src/Data/List/Membership/Setoid/Properties.agda b/src/Data/List/Membership/Setoid/Properties.agda index 1a74ace45f..c06897367b 100644 --- a/src/Data/List/Membership/Setoid/Properties.agda +++ b/src/Data/List/Membership/Setoid/Properties.agda @@ -21,17 +21,17 @@ import Data.List.Relation.Binary.Equality.Setoid as Equality import Data.List.Relation.Unary.Unique.Setoid as Unique open import Data.Nat.Base using (suc; z≤n; s≤s; _≤_; _<_) open import Data.Nat.Properties using (≤-trans; n≤1+n) -open import Data.Product.Base as Prod using (∃; _×_; _,_ ; ∃₂; proj₁; proj₂) +open import Data.Product.Base as Product using (∃; _×_; _,_ ; ∃₂; proj₁; proj₂) open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) open import Data.Sum.Base using (_⊎_; inj₁; inj₂; [_,_]′) open import Function.Base using (_$_; flip; _∘_; _∘′_; id) open import Function.Bundles using (_↔_) open import Level using (Level) open import Relation.Binary.Core using (Rel; _Preserves₂_⟶_⟶_; _Preserves_⟶_) -open import Relation.Binary.Definitions as B hiding (Decidable) +open import Relation.Binary.Definitions as Binary hiding (Decidable) open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) -open import Relation.Unary as U using (Decidable; Pred) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Unary as Unary using (Decidable; Pred) open import Relation.Nullary using (¬_; does; _because_; yes; no) open import Relation.Nullary.Reflects using (invert) open import Relation.Nullary.Negation using (contradiction) @@ -85,11 +85,11 @@ module _ (S : Setoid c ℓ) where ∉×∈⇒≉ : ∀ {x y xs} → All (y ≉_) xs → x ∈ xs → x ≉ y ∉×∈⇒≉ = All.lookupWith λ y≉z x≈z x≈y → y≉z (trans (sym x≈y) x≈z) - unique⇒irrelevant : B.Irrelevant _≈_ → ∀ {xs} → Unique xs → U.Irrelevant (_∈ xs) + unique⇒irrelevant : Binary.Irrelevant _≈_ → ∀ {xs} → Unique xs → Unary.Irrelevant (_∈ xs) unique⇒irrelevant ≈-irr _ (here p) (here q) = - P.cong here (≈-irr p q) + ≡.cong here (≈-irr p q) unique⇒irrelevant ≈-irr (_ ∷ u) (there p) (there q) = - P.cong there (unique⇒irrelevant ≈-irr u p q) + ≡.cong there (unique⇒irrelevant ≈-irr u p q) unique⇒irrelevant ≈-irr (≉s ∷ _) (here p) (there q) = contradiction p (∉×∈⇒≉ ≉s q) unique⇒irrelevant ≈-irr (≉s ∷ _) (there p) (here q) = @@ -130,18 +130,18 @@ module _ (S : Setoid c ℓ) where length-mapWith∈ : ∀ {a} {A : Set a} xs {f : ∀ {x} → x ∈ xs → A} → length (mapWith∈ xs f) ≡ length xs - length-mapWith∈ [] = P.refl - length-mapWith∈ (x ∷ xs) = P.cong suc (length-mapWith∈ xs) + length-mapWith∈ [] = ≡.refl + length-mapWith∈ (x ∷ xs) = ≡.cong suc (length-mapWith∈ xs) mapWith∈-id : ∀ xs → mapWith∈ xs (λ {x} _ → x) ≡ xs - mapWith∈-id [] = P.refl - mapWith∈-id (x ∷ xs) = P.cong (x ∷_) (mapWith∈-id xs) + mapWith∈-id [] = ≡.refl + mapWith∈-id (x ∷ xs) = ≡.cong (x ∷_) (mapWith∈-id xs) map-mapWith∈ : ∀ {a b} {A : Set a} {B : Set b} → ∀ xs (f : ∀ {x} → x ∈ xs → A) (g : A → B) → map g (mapWith∈ xs f) ≡ mapWith∈ xs (g ∘′ f) - map-mapWith∈ [] f g = P.refl - map-mapWith∈ (x ∷ xs) f g = P.cong (_ ∷_) (map-mapWith∈ xs (f ∘ there) g) + map-mapWith∈ [] f g = ≡.refl + map-mapWith∈ (x ∷ xs) f g = ≡.cong (_ ∷_) (map-mapWith∈ xs (f ∘ there) g) ------------------------------------------------------------------------ -- map @@ -164,8 +164,8 @@ module _ (S₁ : Setoid c₁ ℓ₁) (S₂ : Setoid c₂ ℓ₂) where map-∷= : ∀ {f} (f≈ : f Preserves _≈₁_ ⟶ _≈₂_) {xs x v} → (x∈xs : x ∈₁ xs) → map f (x∈xs M₁.∷= v) ≡ ∈-map⁺ f≈ x∈xs M₂.∷= f v - map-∷= f≈ (here x≈y) = P.refl - map-∷= f≈ (there x∈xs) = P.cong (_ ∷_) (map-∷= f≈ x∈xs) + map-∷= f≈ (here x≈y) = ≡.refl + map-∷= f≈ (there x∈xs) = ≡.cong (_ ∷_) (map-∷= f≈ x∈xs) ------------------------------------------------------------------------ -- _++_ @@ -342,15 +342,15 @@ module _ (S : Setoid c ℓ) {P : Pred (Carrier S) p} ∈-filter⁻ : ∀ {v xs} → v ∈ filter P? xs → v ∈ xs × P v ∈-filter⁻ {xs = x ∷ xs} v∈f[x∷xs] with P? x - ... | false because _ = Prod.map there id (∈-filter⁻ v∈f[x∷xs]) + ... | false because _ = Product.map there id (∈-filter⁻ v∈f[x∷xs]) ... | true because [Px] with v∈f[x∷xs] ... | here v≈x = here v≈x , resp (sym v≈x) (invert [Px]) - ... | there v∈fxs = Prod.map there id (∈-filter⁻ v∈fxs) + ... | there v∈fxs = Product.map there id (∈-filter⁻ v∈fxs) ------------------------------------------------------------------------ -- derun and deduplicate -module _ (S : Setoid c ℓ) {R : Rel (Carrier S) ℓ₂} (R? : B.Decidable R) where +module _ (S : Setoid c ℓ) {R : Rel (Carrier S) ℓ₂} (R? : Binary.Decidable R) where open Setoid S using (_≈_) open Membership S using (_∈_) diff --git a/src/Data/List/Nary/NonDependent.agda b/src/Data/List/Nary/NonDependent.agda index f95da6345c..c008efa6c3 100644 --- a/src/Data/List/Nary/NonDependent.agda +++ b/src/Data/List/Nary/NonDependent.agda @@ -10,7 +10,7 @@ module Data.List.Nary.NonDependent where open import Data.Nat.Base using (zero; suc) open import Data.List.Base as List using (List; []; _∷_) -open import Data.Product.Base as Prod using (_,_) +open import Data.Product.Base as Product using (_,_) open import Data.Product.Nary.NonDependent using (Product) open import Function.Base using () open import Function.Nary.NonDependent.Base @@ -37,7 +37,7 @@ zipWith : ∀ n {ls} {as : Sets n ls} {r} {R : Set r} → zipWith 0 f = [] zipWith 1 f xs = List.map f xs zipWith (suc n@(suc _)) f xs ys = - zipWith n (Prod.uncurry f) (List.zipWith _,_ xs ys) + zipWith n (Product.uncurry f) (List.zipWith _,_ xs ys) unzipWith : ∀ n {ls} {as : Sets n ls} {a} {A : Set a} → (A → Product n as) → (List A → Product n (List <$> as)) diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda index 5bd47dc763..0dd6574711 100644 --- a/src/Data/List/Properties.agda +++ b/src/Data/List/Properties.agda @@ -25,9 +25,9 @@ open import Data.Maybe.Base as Maybe using (Maybe; just; nothing) open import Data.Nat.Base open import Data.Nat.Divisibility open import Data.Nat.Properties -open import Data.Product.Base as Prod +open import Data.Product.Base as Product using (_×_; _,_; uncurry; uncurry′; proj₁; proj₂; <_,_>) -import Data.Product.Relation.Unary.All as Prod using (All) +import Data.Product.Relation.Unary.All as Product using (All) open import Data.Sum.Base using (_⊎_; inj₁; inj₂) open import Data.These.Base as These using (These; this; that; these) open import Data.Fin.Properties using (toℕ-cast) @@ -35,8 +35,8 @@ open import Function.Base using (id; _∘_; _∘′_; _∋_; _-⟨_∣; ∣_⟩- open import Function.Definitions using (Injective) open import Level using (Level) open import Relation.Binary.Definitions as B using (DecidableEquality) -import Relation.Binary.Reasoning.Setoid as EqR -open import Relation.Binary.PropositionalEquality as P hiding ([_]) +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning +open import Relation.Binary.PropositionalEquality as ≡ hiding ([_]) open import Relation.Binary.Core using (Rel) open import Relation.Nullary.Reflects using (invert) open import Relation.Nullary using (¬_; Dec; does; _because_; yes; no; contradiction) @@ -370,20 +370,20 @@ module _ (f : A → B → C) where unalignWith-this : unalignWith ((A → These A B) ∋ this) ≗ (_, []) unalignWith-this [] = refl -unalignWith-this (a ∷ as) = cong (Prod.map₁ (a ∷_)) (unalignWith-this as) +unalignWith-this (a ∷ as) = cong (Product.map₁ (a ∷_)) (unalignWith-this as) unalignWith-that : unalignWith ((B → These A B) ∋ that) ≗ ([] ,_) unalignWith-that [] = refl -unalignWith-that (b ∷ bs) = cong (Prod.map₂ (b ∷_)) (unalignWith-that bs) +unalignWith-that (b ∷ bs) = cong (Product.map₂ (b ∷_)) (unalignWith-that bs) module _ {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 = cong (Prod.map₁ (a ∷_)) (unalignWith-cong f≗g cs) - ... | that b | ._ | refl = cong (Prod.map₂ (b ∷_)) (unalignWith-cong f≗g cs) - ... | these a b | ._ | refl = cong (Prod.map (a ∷_) (b ∷_)) (unalignWith-cong f≗g cs) + ... | this a | ._ | refl = cong (Product.map₁ (a ∷_)) (unalignWith-cong f≗g cs) + ... | that b | ._ | refl = cong (Product.map₂ (b ∷_)) (unalignWith-cong f≗g cs) + ... | these a b | ._ | refl = cong (Product.map (a ∷_) (b ∷_)) (unalignWith-cong f≗g cs) module _ (f : C → These A B) where @@ -391,17 +391,17 @@ module _ (f : C → These A B) where unalignWith f (map g ds) ≡ unalignWith (f ∘′ g) ds unalignWith-map g [] = refl unalignWith-map g (d ∷ ds) with f (g d) - ... | this a = cong (Prod.map₁ (a ∷_)) (unalignWith-map g ds) - ... | that b = cong (Prod.map₂ (b ∷_)) (unalignWith-map g ds) - ... | these a b = cong (Prod.map (a ∷_) (b ∷_)) (unalignWith-map g ds) + ... | this a = cong (Product.map₁ (a ∷_)) (unalignWith-map g ds) + ... | that b = cong (Product.map₂ (b ∷_)) (unalignWith-map g ds) + ... | these a b = cong (Product.map (a ∷_) (b ∷_)) (unalignWith-map g ds) map-unalignWith : (g : A → D) (h : B → E) → - Prod.map (map g) (map h) ∘′ unalignWith f ≗ unalignWith (These.map g h ∘′ f) + Product.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 = cong (Prod.map₁ (g a ∷_)) (map-unalignWith g h cs) - ... | that b = cong (Prod.map₂ (h b ∷_)) (map-unalignWith g h cs) - ... | these a b = cong (Prod.map (g a ∷_) (h b ∷_)) (map-unalignWith g h cs) + ... | this a = cong (Product.map₁ (g a ∷_)) (map-unalignWith g h cs) + ... | that b = cong (Product.map₂ (h b ∷_)) (map-unalignWith g h cs) + ... | these a b = cong (Product.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) @@ -417,7 +417,7 @@ module _ (f : C → These A B) where as , [] ∎ unalignWith-alignWith g g∘f≗id (a ∷ as) (b ∷ bs) rewrite g∘f≗id (these a b) = - cong (Prod.map (a ∷_) (b ∷_)) (unalignWith-alignWith g g∘f≗id as bs) + cong (Product.map (a ∷_) (b ∷_)) (unalignWith-alignWith g g∘f≗id as bs) ------------------------------------------------------------------------ -- unzipWith @@ -896,7 +896,7 @@ lookup-iterate f x (suc n) (suc i) = lookup-iterate f (f x) n i 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) = cong (Prod.map (x ∷_) id) (splitAt-defn n xs) +splitAt-defn (suc n) (x ∷ xs) = cong (Product.map (x ∷_) id) (splitAt-defn n xs) ------------------------------------------------------------------------ -- takeWhile, dropWhile, and span @@ -912,7 +912,7 @@ module _ {P : Pred A p} (P? : Decidable P) where span-defn : span P? ≗ < takeWhile P? , dropWhile P? > span-defn [] = refl span-defn (x ∷ xs) with does (P? x) - ... | true = cong (Prod.map (x ∷_) id) (span-defn xs) + ... | true = cong (Product.map (x ∷_) id) (span-defn xs) ... | false = refl ------------------------------------------------------------------------ @@ -1023,15 +1023,15 @@ module _ {P : Pred A p} (P? : Decidable P) where partition-defn : partition P? ≗ < filter P? , filter (∁? P?) > partition-defn [] = refl partition-defn (x ∷ xs) with ih ← partition-defn xs | does (P? x) - ... | true = cong (Prod.map (x ∷_) id) ih - ... | false = cong (Prod.map id (x ∷_)) ih + ... | true = cong (Product.map (x ∷_) id) ih + ... | false = cong (Product.map id (x ∷_)) ih length-partition : ∀ xs → (let (ys , zs) = partition P? xs) → length ys ≤ length xs × length zs ≤ length xs length-partition [] = z≤n , z≤n length-partition (x ∷ xs) with ih ← length-partition xs | does (P? x) - ... | true = Prod.map s≤s m≤n⇒m≤1+n ih - ... | false = Prod.map m≤n⇒m≤1+n s≤s ih + ... | true = Product.map s≤s m≤n⇒m≤1+n ih + ... | false = Product.map m≤n⇒m≤1+n s≤s ih ------------------------------------------------------------------------ -- _ʳ++_ @@ -1181,7 +1181,7 @@ module _ {x y : A} where ∷ʳ-injective : ∀ xs ys → xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys × x ≡ y ∷ʳ-injective [] [] refl = (refl , refl) ∷ʳ-injective (x ∷ xs) (y ∷ ys) eq with refl , eq′ ← ∷-injective eq - = Prod.map (cong (x ∷_)) id (∷ʳ-injective xs ys eq′) + = Product.map (cong (x ∷_)) id (∷ʳ-injective xs ys eq′) ∷ʳ-injective [] (_ ∷ _ ∷ _) () ∷ʳ-injective (_ ∷ _ ∷ _) [] () diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 17af3d8955..389721d4c5 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -15,10 +15,11 @@ open import Data.Empty open import Data.Fin.Base open import Data.List.Base open import Data.List.Effectful using (monad; module Applicative; module MonadProperties) -import Data.List.Properties as LP +import Data.List.Properties as List open import Data.List.Relation.Unary.Any using (Any; here; there) open import Data.List.Relation.Unary.Any.Properties hiding (++-comm) open import Data.List.Membership.Propositional using (_∈_) +open import Data.List.Membership.Propositional.Properties open import Data.List.Relation.Binary.Subset.Propositional.Properties using (⊆-preorder) open import Data.List.Relation.Binary.Permutation.Propositional @@ -39,13 +40,12 @@ open import Level using (Level) open import Relation.Binary.Core using (_⇒_) open import Relation.Binary.Definitions using (Trans) open import Relation.Binary.Bundles using (Preorder; Setoid) -import Relation.Binary.Reasoning.Setoid as EqR -import Relation.Binary.Reasoning.Preorder as PreorderReasoning -open import Relation.Binary.PropositionalEquality as P - using (_≡_; _≢_; _≗_; refl) +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning +import Relation.Binary.Reasoning.Preorder as ≲-Reasoning +open import Relation.Binary.PropositionalEquality as ≡ + using (_≡_; _≢_; _≗_; refl; module ≡-Reasoning) open import Relation.Binary.Reasoning.Syntax open import Relation.Nullary.Negation using (¬_) -open import Data.List.Membership.Propositional.Properties private variable @@ -93,7 +93,7 @@ bag-=⇒ xs≈ys = ↔⇒ xs≈ys -- "Equational" reasoning for _⊆_ along with an additional relatedness module ⊆-Reasoning {A : Set a} where - private module Base = PreorderReasoning (⊆-preorder A) + private module Base = ≲-Reasoning (⊆-preorder A) open Base public hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨; step-∼; step-≲) @@ -140,10 +140,10 @@ module _ {k} {f g : A → B} {xs ys} where helper : ∀ y → x ≡ f y ↔ x ≡ g y helper y = mk↔ₛ′ - (λ x≡fy → P.trans x≡fy ( f≗g y)) - (λ x≡gy → P.trans x≡gy (P.sym $ f≗g y)) - (λ { P.refl → P.trans-symˡ (f≗g y) }) - λ { P.refl → P.trans-symʳ (f≗g y) } + (λ x≡fy → ≡.trans x≡fy ( f≗g y)) + (λ x≡gy → ≡.trans x≡gy (≡.sym $ f≗g y)) + (λ { ≡.refl → ≡.trans-symˡ (f≗g y) }) + λ { ≡.refl → ≡.trans-symʳ (f≗g y) } ------------------------------------------------------------------------ -- _++_ @@ -195,11 +195,11 @@ module _ {k} {A B : Set a} {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 {x} = begin x ∈ (fs ⊛ xs) - ≡⟨ P.cong (x ∈_) (Applicative.unfold-⊛ fs xs) ⟩ + ≡⟨ ≡.cong (x ∈_) (Applicative.unfold-⊛ fs xs) ⟩ x ∈ (fs >>= λ f → xs >>= λ x → pure (f x)) ∼⟨ >>=-cong fs≈gs (λ f → >>=-cong xs≈ys λ x → K-refl) ⟩ x ∈ (gs >>= λ g → ys >>= λ y → pure (g y)) - ≡⟨ P.cong (x ∈_) (Applicative.unfold-⊛ gs ys) ⟨ + ≡⟨ ≡.cong (x ∈_) (Applicative.unfold-⊛ gs ys) ⟨ x ∈ (gs ⊛ ys) ∎ where open Related.EquationalReasoning @@ -232,7 +232,7 @@ commutativeMonoid {a} k A = record ; ∙-cong = ++-cong } ; assoc = λ xs ys zs → - Eq.reflexive (LP.++-assoc xs ys zs) + Eq.reflexive (List.++-assoc xs ys zs) } ; identityˡ = λ xs → K-refl ; comm = λ xs ys → ↔⇒ (++↔++ xs ys) @@ -284,10 +284,10 @@ empty-unique {xs = _ ∷ _} ∷∼[] with ⇒→ ∷∼[] (here refl) (xs₂ >>= pure ∘ f)) ≈⟨ >>=-left-distributive fs ⟩ (fs >>= λ f → xs₁ >>= pure ∘ f) ++ - (fs >>= λ f → xs₂ >>= pure ∘ f) ≡⟨ P.cong₂ _++_ (Applicative.unfold-⊛ fs xs₁) (Applicative.unfold-⊛ fs xs₂) ⟨ + (fs >>= λ f → xs₂ >>= pure ∘ f) ≡⟨ ≡.cong₂ _++_ (Applicative.unfold-⊛ fs xs₁) (Applicative.unfold-⊛ fs xs₂) ⟨ (fs ⊛ xs₁) ++ (fs ⊛ xs₂) ∎ - where open EqR ([ bag ]-Equality B) + where open ≈-Reasoning ([ bag ]-Equality B) private @@ -397,7 +397,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = to (Fin-length-cong xs≈ys) (index-of p) ∎ where - open P.≡-Reasoning + open ≡-Reasoning open Inverse lemma : @@ -423,11 +423,11 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = index-of (Inverse.to xs≈ys q) index-equality-preserved {p = p} {q} xs≈ys eq = index-of (Inverse.to xs≈ys p) ≡⟨ index-of-commutes xs≈ys p ⟩ - Inverse.to (Fin-length-cong xs≈ys) (index-of p) ≡⟨ P.cong (Inverse.to (Fin-length-cong xs≈ys)) eq ⟩ - Inverse.to (Fin-length-cong xs≈ys) (index-of q) ≡⟨ P.sym $ index-of-commutes xs≈ys q ⟩ + Inverse.to (Fin-length-cong xs≈ys) (index-of p) ≡⟨ ≡.cong (Inverse.to (Fin-length-cong xs≈ys)) eq ⟩ + Inverse.to (Fin-length-cong xs≈ys) (index-of q) ≡⟨ ≡.sym $ index-of-commutes xs≈ys q ⟩ index-of (Inverse.to xs≈ys q) ∎ where - open P.≡-Reasoning + open ≡-Reasoning -- The old inspect idiom. @@ -487,17 +487,17 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ∀ b → g (from f) from-hyp (g (to f) to-hyp b) ≡ b g∘g f to-hyp from-hyp b = g∘g′ where - open P.≡-Reasoning + open ≡-Reasoning g∘g′ : g (from f) from-hyp (g (to f) to-hyp b) ≡ b g∘g′ with inspect (to f (inj₂ b)) g∘g′ | inj₂ c , eq₁ with inspect (from f (inj₂ c)) ... | inj₂ b′ , eq₂ = inj₂-injective ( - inj₂ b′ ≡⟨ P.sym eq₂ ⟩ + inj₂ b′ ≡⟨ ≡.sym eq₂ ⟩ from f (inj₂ c) ≡⟨ to-from f eq₁ ⟩ inj₂ b ∎) ... | inj₁ a , eq₂ with - inj₁ a ≡⟨ P.sym eq₂ ⟩ + inj₁ a ≡⟨ ≡.sym eq₂ ⟩ from f (inj₂ c) ≡⟨ to-from f eq₁ ⟩ inj₂ b ∎ ... | () @@ -505,7 +505,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = g∘g′ | inj₁ a , eq₁ | inj₁ a′ , eq₂ = ⊥-elim $ to-hyp eq₁ eq₂ g∘g′ | inj₁ a , eq₁ | inj₂ c , eq₂ with inspect (from f (inj₂ c)) g∘g′ | inj₁ a , eq₁ | inj₂ c , eq₂ | inj₂ b′ , eq₃ with - inj₁ a ≡⟨ P.sym (to-from f eq₂) ⟩ + inj₁ a ≡⟨ ≡.sym (to-from f eq₂) ⟩ from f (inj₂ c) ≡⟨ eq₃ ⟩ inj₂ b′ ∎ ... | () @@ -513,12 +513,12 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = g∘g′ | inj₁ a , eq₁ | inj₂ c , eq₂ | inj₁ a′ , eq₃ | inj₁ a″ , eq₄ = ⊥-elim $ from-hyp eq₃ eq₄ g∘g′ | inj₁ a , eq₁ | inj₂ c , eq₂ | inj₁ a′ , eq₃ | inj₂ b′ , eq₄ = inj₂-injective ( let lemma = - inj₁ a′ ≡⟨ P.sym eq₃ ⟩ + inj₁ a′ ≡⟨ ≡.sym eq₃ ⟩ from f (inj₂ c) ≡⟨ to-from f eq₂ ⟩ inj₁ a ∎ in - inj₂ b′ ≡⟨ P.sym eq₄ ⟩ - from f (inj₁ a′) ≡⟨ P.cong (from f ∘ inj₁) $ inj₁-injective lemma ⟩ + inj₂ b′ ≡⟨ ≡.sym eq₄ ⟩ + from f (inj₁ a′) ≡⟨ ≡.cong (from f ∘ inj₁) $ inj₁-injective lemma ⟩ from f (inj₁ a) ≡⟨ to-from f eq₁ ⟩ inj₂ b ∎) @@ -541,21 +541,21 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ with zero ≡⟨⟩ index-of {xs = x ∷ xs} (here p) ≡⟨⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₁ p) ≡⟨ P.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ P.sym $ + index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₁ p) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ ≡.sym $ to-from (∼→⊎↔⊎ inv) {x = inj₁ p} hyp₂ ⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) $ (from (∼→⊎↔⊎ inv) $ inj₁ q)) ≡⟨ P.cong index-of $ + index-of {xs = x ∷ xs} (to (∷↔ _) $ (from (∼→⊎↔⊎ inv) $ inj₁ q)) ≡⟨ ≡.cong index-of $ strictlyInverseˡ (∷↔ _) (from inv (here q)) ⟩ index-of {xs = x ∷ xs} (to (SK-sym inv) $ here q) ≡⟨ index-equality-preserved (SK-sym inv) refl ⟩ - index-of {xs = x ∷ xs} (to (SK-sym inv) $ here p) ≡⟨ P.cong index-of $ P.sym $ + index-of {xs = x ∷ xs} (to (SK-sym inv) $ here p) ≡⟨ ≡.cong index-of $ ≡.sym $ strictlyInverseˡ (∷↔ _) (from inv (here p)) ⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) (from (∼→⊎↔⊎ inv) $ inj₁ p)) ≡⟨ P.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ + index-of {xs = x ∷ xs} (to (∷↔ _) (from (∼→⊎↔⊎ inv) $ inj₁ p)) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ to-from (∼→⊎↔⊎ inv) {x = inj₂ z∈xs} hyp₁ ⟩ index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₂ z∈xs) ≡⟨⟩ index-of {xs = x ∷ xs} (there z∈xs) ≡⟨⟩ suc (index-of {xs = xs} z∈xs) ∎ where open Inverse - open P.≡-Reasoning + open ≡-Reasoning ... | () ------------------------------------------------------------------------ @@ -573,10 +573,10 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷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) = P.cong there (from∘to p v∈xs) + 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)) = P.cong (there ∘ there) (from∘to p v∈xs) + 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 @@ -588,7 +588,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A} ∼bag⇒↭ {A = A} {[]} eq with empty-unique (↔-sym eq) ... | refl = refl -∼bag⇒↭ {A = A} {x ∷ xs} eq with ∈-∃++ (Inverse.to (eq {x}) (here P.refl)) +∼bag⇒↭ {A = A} {x ∷ xs} eq with ∈-∃++ (Inverse.to (eq {x}) (here ≡.refl)) ... | zs₁ , zs₂ , p rewrite p = begin x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩ x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩ diff --git a/src/Data/List/Relation/Binary/Infix/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Infix/Heterogeneous/Properties.agda index 9af0162ded..7e85ae8661 100644 --- a/src/Data/List/Relation/Binary/Infix/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Infix/Heterogeneous/Properties.agda @@ -13,7 +13,7 @@ open import Data.Bool.Base using (true; false) open import Data.Empty using (⊥-elim) open import Data.List.Base as List using (List; []; _∷_; length; map; filter; replicate) open import Data.Nat.Base using (zero; suc; _≤_) -import Data.Nat.Properties as ℕₚ +import Data.Nat.Properties as ℕ open import Data.Sum.Base using (_⊎_; inj₁; inj₂; [_,_]′) open import Function.Base using (case_of_; _$′_) @@ -27,7 +27,7 @@ open import Data.List.Relation.Binary.Pointwise.Base as Pointwise using (Pointwi open import Data.List.Relation.Binary.Infix.Heterogeneous open import Data.List.Relation.Binary.Prefix.Heterogeneous as Prefix using (Prefix; []; _∷_) -import Data.List.Relation.Binary.Prefix.Heterogeneous.Properties as Prefixₚ +import Data.List.Relation.Binary.Prefix.Heterogeneous.Properties as Prefix open import Data.List.Relation.Binary.Suffix.Heterogeneous as Suffix using (Suffix; here; there) @@ -43,7 +43,7 @@ private -- Conversion functions fromPointwise : ∀ {as bs} → Pointwise R as bs → Infix R as bs -fromPointwise pw = here (Prefixₚ.fromPointwise pw) +fromPointwise pw = here (Prefix.fromPointwise pw) fromSuffix : ∀ {as bs} → Suffix R as bs → Infix R as bs fromSuffix (here pw) = fromPointwise pw @@ -52,11 +52,11 @@ fromSuffix (there p) = there (fromSuffix p) module _ {c t} {C : Set c} {T : REL A C t} where fromPrefixSuffix : Trans R S T → Trans (Prefix R) (Suffix S) (Infix T) - fromPrefixSuffix tr p (here q) = here (Prefixₚ.trans tr p (Prefixₚ.fromPointwise q)) + fromPrefixSuffix tr p (here q) = here (Prefix.trans tr p (Prefix.fromPointwise q)) fromPrefixSuffix tr p (there q) = there (fromPrefixSuffix tr p q) fromSuffixPrefix : Trans R S T → Trans (Suffix R) (Prefix S) (Infix T) - fromSuffixPrefix tr (here p) q = here (Prefixₚ.trans tr (Prefixₚ.fromPointwise p) q) + fromSuffixPrefix tr (here p) q = here (Prefix.trans tr (Prefix.fromPointwise p) q) fromSuffixPrefix tr (there p) (_ ∷ q) = there (fromSuffixPrefix tr p q) ∷⁻ : ∀ {as b bs} → Infix R as (b ∷ bs) → Prefix R as (b ∷ bs) ⊎ Infix R as bs @@ -67,8 +67,8 @@ module _ {c t} {C : Set c} {T : REL A C t} where -- length length-mono : ∀ {as bs} → Infix R as bs → length as ≤ length bs -length-mono (here pref) = Prefixₚ.length-mono pref -length-mono (there p) = ℕₚ.m≤n⇒m≤1+n (length-mono p) +length-mono (here pref) = Prefix.length-mono pref +length-mono (there p) = ℕ.m≤n⇒m≤1+n (length-mono p) ------------------------------------------------------------------------ -- As an order @@ -76,11 +76,11 @@ length-mono (there p) = ℕₚ.m≤n⇒m≤1+n (length-mono p) module _ {c t} {C : Set c} {T : REL A C t} where Prefix-Infix-trans : Trans R S T → Trans (Prefix R) (Infix S) (Infix T) - Prefix-Infix-trans tr p (here q) = here (Prefixₚ.trans tr p q) + Prefix-Infix-trans tr p (here q) = here (Prefix.trans tr p q) Prefix-Infix-trans tr p (there q) = there (Prefix-Infix-trans tr p q) Infix-Prefix-trans : Trans R S T → Trans (Infix R) (Prefix S) (Infix T) - Infix-Prefix-trans tr (here p) q = here (Prefixₚ.trans tr p q) + Infix-Prefix-trans tr (here p) q = here (Prefix.trans tr p q) Infix-Prefix-trans tr (there p) (_ ∷ q) = there (Infix-Prefix-trans tr p q) Suffix-Infix-trans : Trans R S T → Trans (Suffix R) (Infix S) (Infix T) @@ -88,7 +88,7 @@ module _ {c t} {C : Set c} {T : REL A C t} where Suffix-Infix-trans tr p (there q) = there (Suffix-Infix-trans tr p q) Infix-Suffix-trans : Trans R S T → Trans (Infix R) (Suffix S) (Infix T) - Infix-Suffix-trans tr p (here q) = Infix-Prefix-trans tr p (Prefixₚ.fromPointwise q) + Infix-Suffix-trans tr p (here q) = Infix-Prefix-trans tr p (Prefix.fromPointwise q) Infix-Suffix-trans tr p (there q) = there (Infix-Suffix-trans tr p q) trans : Trans R S T → Trans (Infix R) (Infix S) (Infix T) @@ -96,22 +96,22 @@ module _ {c t} {C : Set c} {T : REL A C t} where trans tr p (there q) = there (trans tr p q) antisym : Antisym R S T → Antisym (Infix R) (Infix S) (Pointwise T) - antisym asym (here p) (here q) = Prefixₚ.antisym asym p q + antisym asym (here p) (here q) = Prefix.antisym asym p q antisym asym {i = a ∷ as} {j = bs} p@(here _) (there q) - = ⊥-elim $′ ℕₚ.<-irrefl refl $′ begin-strict + = ⊥-elim $′ ℕ.<-irrefl refl $′ begin-strict length as <⟨ length-mono p ⟩ length bs ≤⟨ length-mono q ⟩ - length as ∎ where open ℕₚ.≤-Reasoning + length as ∎ where open ℕ.≤-Reasoning antisym asym {i = as} {j = b ∷ bs} (there p) q@(here _) - = ⊥-elim $′ ℕₚ.<-irrefl refl $′ begin-strict + = ⊥-elim $′ ℕ.<-irrefl refl $′ begin-strict length bs <⟨ length-mono q ⟩ length as ≤⟨ length-mono p ⟩ - length bs ∎ where open ℕₚ.≤-Reasoning + length bs ∎ where open ℕ.≤-Reasoning antisym asym {i = a ∷ as} {j = b ∷ bs} (there p) (there q) - = ⊥-elim $′ ℕₚ.<-irrefl refl $′ begin-strict + = ⊥-elim $′ ℕ.<-irrefl refl $′ begin-strict length as <⟨ length-mono p ⟩ length bs <⟨ length-mono q ⟩ - length as ∎ where open ℕₚ.≤-Reasoning + length as ∎ where open ℕ.≤-Reasoning ------------------------------------------------------------------------ -- map @@ -121,14 +121,14 @@ module _ {c d r} {C : Set c} {D : Set d} {R : REL C D r} where map⁺ : ∀ {as bs} (f : A → C) (g : B → D) → Infix (λ a b → R (f a) (g b)) as bs → Infix R (List.map f as) (List.map g bs) - map⁺ f g (here p) = here (Prefixₚ.map⁺ f g p) + map⁺ f g (here p) = here (Prefix.map⁺ f g p) map⁺ f g (there p) = there (map⁺ f g p) map⁻ : ∀ {as bs} (f : A → C) (g : B → D) → Infix R (List.map f as) (List.map g bs) → Infix (λ a b → R (f a) (g b)) as bs - map⁻ {bs = []} f g (here p) = here (Prefixₚ.map⁻ f g p) - map⁻ {bs = b ∷ bs} f g (here p) = here (Prefixₚ.map⁻ f g p) + map⁻ {bs = []} f g (here p) = here (Prefix.map⁻ f g p) + map⁻ {bs = b ∷ bs} f g (here p) = here (Prefix.map⁻ f g p) map⁻ {bs = b ∷ bs} f g (there p) = there (map⁻ f g p) ------------------------------------------------------------------------ @@ -139,7 +139,7 @@ module _ {p q} {P : Pred A p} {Q : Pred B q} (P? : U.Decidable P) (Q? : U.Decida where filter⁺ : ∀ {as bs} → Infix R as bs → Infix R (filter P? as) (filter Q? bs) - filter⁺ (here p) = here (Prefixₚ.filter⁺ P? Q? (λ _ → P⇒Q) (λ _ → Q⇒P) p) + filter⁺ (here p) = here (Prefix.filter⁺ P? Q? (λ _ → P⇒Q) (λ _ → Q⇒P) p) filter⁺ {bs = b ∷ bs} (there p) with does (Q? b) ... | true = there (filter⁺ p) ... | false = filter⁺ p @@ -149,12 +149,12 @@ module _ {p q} {P : Pred A p} {Q : Pred B q} (P? : U.Decidable P) (Q? : U.Decida replicate⁺ : ∀ {m n a b} → m ≤ n → R a b → Infix R (replicate m a) (replicate n b) -replicate⁺ m≤n r = here (Prefixₚ.replicate⁺ m≤n r) +replicate⁺ m≤n r = here (Prefix.replicate⁺ m≤n r) replicate⁻ : ∀ {m n a b} → m ≢ 0 → Infix R (replicate m a) (replicate n b) → R a b -replicate⁻ {m = m} {n = zero} m≢0 (here p) = Prefixₚ.replicate⁻ m≢0 p -replicate⁻ {m = m} {n = suc n} m≢0 (here p) = Prefixₚ.replicate⁻ m≢0 p +replicate⁻ {m = m} {n = zero} m≢0 (here p) = Prefix.replicate⁻ m≢0 p +replicate⁻ {m = m} {n = suc n} m≢0 (here p) = Prefix.replicate⁻ m≢0 p replicate⁻ {m = m} {n = suc n} m≢0 (there p) = replicate⁻ m≢0 p ------------------------------------------------------------------------ @@ -165,4 +165,4 @@ infix? R? [] [] = yes (here []) infix? R? (a ∷ as) [] = no (λ where (here ())) infix? R? as bbs@(_ ∷ bs) = map′ [ here , there ]′ ∷⁻ - (Prefixₚ.prefix? R? as bbs ⊎-dec infix? R? as bs) + (Prefix.prefix? R? as bbs ⊎-dec infix? R? as bs) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index edcf62b6da..4409e78cc4 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -24,7 +24,7 @@ open import Data.List.Relation.Binary.Equality.Setoid S open import Data.Nat.Base using (ℕ; zero; suc; _+_) open import Level using (_⊔_) open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) -import Relation.Binary.Reasoning.Setoid as SetoidReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning private module Eq = Setoid S @@ -88,7 +88,7 @@ steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs module PermutationReasoning where - private module Base = SetoidReasoning ↭-setoid + private module Base = ≈-Reasoning ↭-setoid open Base public hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨) diff --git a/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda index 202bff89b0..776d065423 100644 --- a/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda @@ -19,7 +19,7 @@ open import Data.List.Relation.Binary.Pointwise.Base using (Pointwise; []; _∷_ open import Data.List.Relation.Binary.Prefix.Heterogeneous as Prefix hiding (PrefixView; _++_) open import Data.Nat.Base using (ℕ; zero; suc; _≤_; z≤n; s≤s) open import Data.Nat.Properties using (suc-injective) -open import Data.Product.Base as Prod using (_×_; _,_; proj₁; proj₂; uncurry) +open import Data.Product.Base as Product using (_×_; _,_; proj₁; proj₂; uncurry) open import Function.Base open import Relation.Nullary.Negation using (¬_) diff --git a/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Properties.agda index b11e011d76..3184ca4420 100644 --- a/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Properties.agda @@ -15,15 +15,15 @@ open import Data.Empty open import Data.List.Relation.Unary.All using (Null; []; _∷_) open import Data.List.Relation.Unary.Any using (Any; here; there) open import Data.List.Base as List hiding (map; _∷ʳ_) -import Data.List.Properties as Lₚ +import Data.List.Properties as List open import Data.List.Relation.Unary.Any.Properties using (here-injective; there-injective) open import Data.List.Relation.Binary.Pointwise as Pw using (Pointwise; []; _∷_) open import Data.List.Relation.Binary.Sublist.Heterogeneous -open import Data.Maybe.Relation.Unary.All as MAll using (nothing; just) +open import Data.Maybe.Relation.Unary.All as Maybe using (nothing; just) open import Data.Nat.Base using (ℕ; _≤_; _≥_); open ℕ; open _≤_ -import Data.Nat.Properties as ℕₚ +import Data.Nat.Properties as ℕ open import Data.Product.Base using (∃₂; _×_; _,_; <_,_>; proj₂; uncurry) open import Function.Base @@ -62,7 +62,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where length-mono-≤ : ∀ {as bs} → Sublist R as bs → length as ≤ length bs length-mono-≤ [] = z≤n - length-mono-≤ (y ∷ʳ rs) = ℕₚ.m≤n⇒m≤1+n (length-mono-≤ rs) + length-mono-≤ (y ∷ʳ rs) = ℕ.m≤n⇒m≤1+n (length-mono-≤ rs) length-mono-≤ (r ∷ rs) = s≤s (length-mono-≤ rs) ------------------------------------------------------------------------ @@ -75,9 +75,9 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where toPointwise : ∀ {as bs} → length as ≡ length bs → Sublist R as bs → Pointwise R as bs toPointwise {bs = []} eq [] = [] - toPointwise {bs = b ∷ bs} eq (r ∷ rs) = r ∷ toPointwise (ℕₚ.suc-injective eq) rs + toPointwise {bs = b ∷ bs} eq (r ∷ rs) = r ∷ toPointwise (ℕ.suc-injective eq) rs toPointwise {bs = b ∷ bs} eq (b ∷ʳ rs) = - ⊥-elim $ ℕₚ.<-irrefl eq (s≤s (length-mono-≤ rs)) + ⊥-elim $ ℕ.<-irrefl eq (s≤s (length-mono-≤ rs)) ------------------------------------------------------------------------ -- Various functions' outputs are sublists @@ -90,9 +90,9 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where tail-Sublist : ∀ {as bs} → Sublist R as bs → - MAll.All (λ as → Sublist R as bs) (tail as) + Maybe.All (λ as → Sublist R as bs) (tail as) tail-Sublist [] = nothing - tail-Sublist (b ∷ʳ ps) = MAll.map (b ∷ʳ_) (tail-Sublist ps) + tail-Sublist (b ∷ʳ ps) = Maybe.map (b ∷ʳ_) (tail-Sublist ps) tail-Sublist (p ∷ ps) = just (_ ∷ʳ ps) take-Sublist : ∀ {as bs} n → Sublist R as bs → Sublist R (take n as) bs @@ -190,7 +190,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where ++⁻ : ∀ {as bs cs ds} → length as ≡ length bs → Sublist R (as ++ cs) (bs ++ ds) → Sublist R cs ds ++⁻ {[]} {[]} eq rs = rs - ++⁻ {a ∷ as} {b ∷ bs} eq rs = ++⁻ (ℕₚ.suc-injective eq) (∷⁻ rs) + ++⁻ {a ∷ as} {b ∷ bs} eq rs = ++⁻ (ℕ.suc-injective eq) (∷⁻ rs) ++ˡ : ∀ {as bs} (cs : List B) → Sublist R as bs → Sublist R as (cs ++ bs) ++ˡ zs = ++⁺ (minimum zs) @@ -223,7 +223,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where Sublist R (drop m as) (drop n bs) drop⁺ {m} z≤n rs = drop-Sublist m rs drop⁺ (s≤s m≥n) [] = [] - drop⁺ (s≤s m≥n) (y ∷ʳ rs) = drop⁺ (ℕₚ.m≤n⇒m≤1+n m≥n) rs + drop⁺ (s≤s m≥n) (y ∷ʳ rs) = drop⁺ (ℕ.m≤n⇒m≤1+n m≥n) rs drop⁺ (s≤s m≥n) (r ∷ rs) = drop⁺ m≥n rs drop⁺-≥ : ∀ {m n as bs} → m ≥ n → Pointwise R as bs → @@ -232,7 +232,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where drop⁺-⊆ : ∀ {as bs} m → Sublist R as bs → Sublist R (drop m as) (drop m bs) - drop⁺-⊆ m = drop⁺ (ℕₚ.≤-refl {m}) + drop⁺-⊆ m = drop⁺ (ℕ.≤-refl {m}) module _ {a b r p q} {A : Set a} {B : Set b} {R : REL A B r} {P : Pred A p} {Q : Pred B q} @@ -308,7 +308,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where reverse⁻ : ∀ {as bs} → Sublist R (reverse as) (reverse bs) → Sublist R as bs reverse⁻ {as} {bs} p = cast (reverse⁺ p) where - cast = P.subst₂ (Sublist R) (Lₚ.reverse-involutive as) (Lₚ.reverse-involutive bs) + cast = P.subst₂ (Sublist R) (List.reverse-involutive as) (List.reverse-involutive bs) ------------------------------------------------------------------------ -- Inversion lemmas @@ -385,25 +385,25 @@ module Antisymmetry {R : REL A B r} {S : REL B A s} {E : REL A B e} (rs⇒e : Antisym R S E) where - open ℕₚ.≤-Reasoning + open ℕ.≤-Reasoning antisym : Antisym (Sublist R) (Sublist S) (Pointwise E) antisym [] [] = [] antisym (r ∷ rs) (s ∷ ss) = rs⇒e r s ∷ antisym rs ss -- impossible cases antisym (_∷ʳ_ {xs} {ys₁} y rs) (_∷ʳ_ {ys₂} {zs} z ss) = - ⊥-elim $ ℕₚ.<-irrefl P.refl $ begin + ⊥-elim $ ℕ.<-irrefl P.refl $ begin length (y ∷ ys₁) ≤⟨ length-mono-≤ ss ⟩ - length zs ≤⟨ ℕₚ.n≤1+n (length zs) ⟩ + length zs ≤⟨ ℕ.n≤1+n (length zs) ⟩ length (z ∷ zs) ≤⟨ length-mono-≤ rs ⟩ length ys₁ ∎ antisym (_∷ʳ_ {xs} {ys₁} y rs) (_∷_ {y} {ys₂} {z} {zs} s ss) = - ⊥-elim $ ℕₚ.<-irrefl P.refl $ begin + ⊥-elim $ ℕ.<-irrefl P.refl $ begin length (z ∷ zs) ≤⟨ length-mono-≤ rs ⟩ length ys₁ ≤⟨ length-mono-≤ ss ⟩ length zs ∎ antisym (_∷_ {x} {xs} {y} {ys₁} r rs) (_∷ʳ_ {ys₂} {zs} z ss) = - ⊥-elim $ ℕₚ.<-irrefl P.refl $ begin + ⊥-elim $ ℕ.<-irrefl P.refl $ begin length (y ∷ ys₁) ≤⟨ length-mono-≤ ss ⟩ length xs ≤⟨ length-mono-≤ rs ⟩ length ys₁ ∎ diff --git a/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Solver.agda b/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Solver.agda index e7b95805be..1813568d5b 100644 --- a/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Solver.agda +++ b/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Solver.agda @@ -21,8 +21,8 @@ module Data.List.Relation.Binary.Sublist.Heterogeneous.Solver open import Level using (_⊔_) open import Data.Fin as Fin -open import Data.Maybe.Base as M -open import Data.Nat.Base as Nat using (ℕ) +open import Data.Maybe.Base as Maybe +open import Data.Nat.Base as ℕ using (ℕ) open import Data.Product.Base using (Σ-syntax; _,_) open import Data.Vec.Base as Vec using (Vec ; lookup) open import Data.List.Base hiding (lookup) @@ -124,8 +124,8 @@ private -- Solver for items solveI : ∀ {n} (a b : Item n) → Maybe (a ⊆I b) -solveI (var k) (var l) = M.map var $ decToMaybe (k Fin.≟ l) -solveI (val a) (val b) = M.map val $ decToMaybe (R? a b) +solveI (var k) (var l) = Maybe.map var $ decToMaybe (k Fin.≟ l) +solveI (val a) (val b) = Maybe.map val $ decToMaybe (R? a b) solveI _ _ = nothing -- Solver for linearised expressions @@ -135,8 +135,8 @@ solveR [] e = just (λ ρ → minimum _) solveR d [] = nothing -- actual work solveR (a ∷ d) (b ∷ e) with solveI a b -... | just it = M.map (keep-it it d e) (solveR d e) -... | nothing = M.map (skip-it b (a ∷ d) e) (solveR (a ∷ d) e) +... | just it = Maybe.map (keep-it it d e) (solveR d e) +... | nothing = Maybe.map (skip-it b (a ∷ d) e) (solveR (a ∷ d) e) -- Coming back to ASTs thanks to flatten diff --git a/src/Data/List/Relation/Binary/Sublist/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Sublist/Setoid/Properties.agda index 19fe765456..cf12328323 100644 --- a/src/Data/List/Relation/Binary/Sublist/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Sublist/Setoid/Properties.agda @@ -16,7 +16,7 @@ open import Data.List.Base hiding (_∷ʳ_) open import Data.List.Relation.Unary.Any using (Any) import Data.Maybe.Relation.Unary.All as Maybe open import Data.Nat.Base using (_≤_; _≥_) -import Data.Nat.Properties as ℕₚ +import Data.Nat.Properties as ℕ open import Data.Product.Base using (∃; _,_; proj₂) open import Function.Base open import Function.Bundles using (_⇔_; _⤖_) @@ -159,7 +159,7 @@ module _ {m n} {xs : List A} where module _ {xs ys : List A} where drop⁺-⊆ : ∀ n → xs ⊆ ys → drop n xs ⊆ drop n ys - drop⁺-⊆ n xs⊆ys = drop⁺ {n} ℕₚ.≤-refl xs⊆ys + drop⁺-⊆ n xs⊆ys = drop⁺ {n} ℕ.≤-refl xs⊆ys ------------------------------------------------------------------------ -- takeWhile / dropWhile diff --git a/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda index 71313d087e..3e91679bb2 100644 --- a/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda @@ -21,12 +21,12 @@ open import Data.List.Effectful open import Data.List.Relation.Unary.Any using (Any) open import Data.List.Membership.Propositional open import Data.List.Membership.Propositional.Properties -import Data.List.Relation.Binary.Subset.Setoid.Properties as Setoidₚ +import Data.List.Relation.Binary.Subset.Setoid.Properties as Subset open import Data.List.Relation.Binary.Subset.Propositional open import Data.List.Relation.Binary.Permutation.Propositional import Data.List.Relation.Binary.Permutation.Propositional.Properties as Permutation open import Data.Nat using (ℕ; _≤_) -import Data.Product.Base as Prod +import Data.Product.Base as Product import Data.Sum.Base as Sum open import Effect.Monad open import Function.Base using (_∘_; _∘′_; id; _$_) @@ -38,7 +38,7 @@ open import Relation.Binary.Core using (_⇒_) open import Relation.Binary.Bundles using (Preorder) open import Relation.Binary.PropositionalEquality using (_≡_; _≗_; isEquivalence; subst; resp; refl; setoid; module ≡-Reasoning) -import Relation.Binary.Reasoning.Preorder as PreorderReasoning +import Relation.Binary.Reasoning.Preorder as ≲-Reasoning private open module ListMonad {ℓ} = RawMonad (monad {ℓ = ℓ}) @@ -79,7 +79,7 @@ module _ (A : Set a) where ------------------------------------------------------------------------ -- Relational properties with _↭_ (permutation) ------------------------------------------------------------------------ --- See issue #1354 for why these proofs can't be taken from `Setoidₚ` +-- See issue #1354 for why these proofs can't be taken from `Subset` ⊆-reflexive-↭ : _↭_ {A = A} ⇒ _⊆_ ⊆-reflexive-↭ xs↭ys = Permutation.∈-resp-↭ xs↭ys @@ -109,7 +109,7 @@ module _ (A : Set a) where ------------------------------------------------------------------------ module ⊆-Reasoning (A : Set a) where - open Setoidₚ.⊆-Reasoning (setoid A) public + open Subset.⊆-Reasoning (setoid A) public hiding (step-≋; step-≋˘) ------------------------------------------------------------------------ @@ -117,10 +117,10 @@ module ⊆-Reasoning (A : Set a) where ------------------------------------------------------------------------ Any-resp-⊆ : ∀ {P : Pred A p} → (Any P) Respects _⊆_ -Any-resp-⊆ = Setoidₚ.Any-resp-⊆ (setoid _) (subst _) +Any-resp-⊆ = Subset.Any-resp-⊆ (setoid _) (subst _) All-resp-⊇ : ∀ {P : Pred A p} → (All P) Respects _⊇_ -All-resp-⊇ = Setoidₚ.All-resp-⊇ (setoid _) (subst _) +All-resp-⊇ = Subset.All-resp-⊇ (setoid _) (subst _) ------------------------------------------------------------------------ -- Properties relating _⊆_ to various list functions @@ -130,38 +130,38 @@ All-resp-⊇ = Setoidₚ.All-resp-⊇ (setoid _) (subst _) map⁺ : ∀ (f : A → B) → xs ⊆ ys → map f xs ⊆ map f ys map⁺ f xs⊆ys = Inverse.to (map-∈↔ f) ∘ - Prod.map₂ (Prod.map₁ xs⊆ys) ∘ + Product.map₂ (Product.map₁ xs⊆ys) ∘ Inverse.from (map-∈↔ f) ------------------------------------------------------------------------ -- ∷ xs⊆x∷xs : ∀ (xs : List A) x → xs ⊆ x ∷ xs -xs⊆x∷xs = Setoidₚ.xs⊆x∷xs (setoid _) +xs⊆x∷xs = Subset.xs⊆x∷xs (setoid _) ∷⁺ʳ : ∀ x → xs ⊆ ys → x ∷ xs ⊆ x ∷ ys -∷⁺ʳ = Setoidₚ.∷⁺ʳ (setoid _) +∷⁺ʳ = Subset.∷⁺ʳ (setoid _) ∈-∷⁺ʳ : ∀ {x} → x ∈ ys → xs ⊆ ys → x ∷ xs ⊆ ys -∈-∷⁺ʳ = Setoidₚ.∈-∷⁺ʳ (setoid _) +∈-∷⁺ʳ = Subset.∈-∷⁺ʳ (setoid _) ------------------------------------------------------------------------ -- _++_ xs⊆xs++ys : ∀ (xs ys : List A) → xs ⊆ xs ++ ys -xs⊆xs++ys = Setoidₚ.xs⊆xs++ys (setoid _) +xs⊆xs++ys = Subset.xs⊆xs++ys (setoid _) xs⊆ys++xs : ∀ (xs ys : List A) → xs ⊆ ys ++ xs -xs⊆ys++xs = Setoidₚ.xs⊆ys++xs (setoid _) +xs⊆ys++xs = Subset.xs⊆ys++xs (setoid _) ++⁺ʳ : ∀ zs → xs ⊆ ys → zs ++ xs ⊆ zs ++ ys -++⁺ʳ = Setoidₚ.++⁺ʳ (setoid _) +++⁺ʳ = Subset.++⁺ʳ (setoid _) ++⁺ˡ : ∀ zs → xs ⊆ ys → xs ++ zs ⊆ ys ++ zs -++⁺ˡ = Setoidₚ.++⁺ˡ (setoid _) +++⁺ˡ = Subset.++⁺ˡ (setoid _) ++⁺ : ws ⊆ xs → ys ⊆ zs → ws ++ ys ⊆ xs ++ zs -++⁺ = Setoidₚ.++⁺ (setoid _) +++⁺ = Subset.++⁺ (setoid _) ------------------------------------------------------------------------ -- concat @@ -171,14 +171,14 @@ module _ {xss yss : List (List A)} where concat⁺ : xss ⊆ yss → concat xss ⊆ concat yss concat⁺ xss⊆yss = Inverse.to concat-∈↔ ∘ - Prod.map₂ (Prod.map₂ xss⊆yss) ∘ + Product.map₂ (Product.map₂ xss⊆yss) ∘ Inverse.from concat-∈↔ ------------------------------------------------------------------------ -- applyUpTo applyUpTo⁺ : ∀ (f : ℕ → A) {m n} → m ≤ n → applyUpTo f m ⊆ applyUpTo f n -applyUpTo⁺ = Setoidₚ.applyUpTo⁺ (setoid _) +applyUpTo⁺ = Subset.applyUpTo⁺ (setoid _) ------------------------------------------------------------------------ -- _>>=_ @@ -188,7 +188,7 @@ module _ {A B : Set a} (f g : A → List B) where >>=⁺ : xs ⊆ ys → (∀ {x} → f x ⊆ g x) → (xs >>= f) ⊆ (ys >>= g) >>=⁺ xs⊆ys f⊆g = Inverse.to >>=-∈↔ ∘ - Prod.map₂ (Prod.map xs⊆ys f⊆g) ∘ + Product.map₂ (Product.map xs⊆ys f⊆g) ∘ Inverse.from >>=-∈↔ ------------------------------------------------------------------------ @@ -199,7 +199,7 @@ module _ {A B : Set a} {fs gs : List (A → B)} where ⊛⁺ : fs ⊆ gs → xs ⊆ ys → (fs ⊛ xs) ⊆ (gs ⊛ ys) ⊛⁺ fs⊆gs xs⊆ys = (Inverse.to $ ⊛-∈↔ gs) ∘ - Prod.map₂ (Prod.map₂ (Prod.map fs⊆gs (Prod.map₁ xs⊆ys))) ∘ + Product.map₂ (Product.map₂ (Product.map fs⊆gs (Product.map₁ xs⊆ys))) ∘ (Inverse.from $ ⊛-∈↔ fs) ------------------------------------------------------------------------ @@ -210,7 +210,7 @@ module _ {A B : Set a} {ws xs : List A} {ys zs : List B} where ⊗⁺ : ws ⊆ xs → ys ⊆ zs → (ws ⊗ ys) ⊆ (xs ⊗ zs) ⊗⁺ ws⊆xs ys⊆zs = Inverse.to ⊗-∈↔ ∘ - Prod.map ws⊆xs ys⊆zs ∘ + Product.map ws⊆xs ys⊆zs ∘ Inverse.from ⊗-∈↔ ------------------------------------------------------------------------ @@ -235,7 +235,7 @@ module _ {xs : List A} {f : ∀ {x} → x ∈ xs → B} mapWith∈ xs f ⊆ mapWith∈ ys g mapWith∈⁺ xs⊆ys f≈g {x} = Inverse.to Any.mapWith∈↔ ∘ - Prod.map₂ (Prod.map xs⊆ys (λ {x∈xs} x≡fx∈xs → begin + Product.map₂ (Product.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) ∎)) ∘ @@ -248,12 +248,12 @@ module _ {xs : List A} {f : ∀ {x} → x ∈ xs → B} module _ {P : Pred A p} (P? : Decidable P) where filter-⊆ : ∀ xs → filter P? xs ⊆ xs - filter-⊆ = Setoidₚ.filter-⊆ (setoid A) P? + filter-⊆ = Subset.filter-⊆ (setoid A) P? module _ {Q : Pred A q} (Q? : Decidable Q) where filter⁺′ : P ⋐ Q → ∀ {xs ys} → xs ⊆ ys → filter P? xs ⊆ filter Q? ys - filter⁺′ = Setoidₚ.filter⁺′ (setoid A) P? (resp P) Q? (resp Q) + filter⁺′ = Subset.filter⁺′ (setoid A) P? (resp P) Q? (resp Q) ------------------------------------------------------------------------ diff --git a/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda index 0c344758b7..bc1a693a2f 100644 --- a/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda @@ -31,7 +31,7 @@ open import Relation.Binary.Definitions using (Reflexive; Transitive; _Respectsʳ_; _Respectsˡ_; _Respects_) open import Relation.Binary.Bundles using (Setoid; Preorder) open import Relation.Binary.Structures using (IsPreorder) -import Relation.Binary.Reasoning.Preorder as PreorderReasoning +import Relation.Binary.Reasoning.Preorder as ≲-Reasoning open import Relation.Binary.Reasoning.Syntax open Setoid using (Carrier) @@ -115,7 +115,7 @@ module _ (S : Setoid a ℓ) where module ⊆-Reasoning (S : Setoid a ℓ) where open Membership S using (_∈_) - private module Base = PreorderReasoning (⊆-preorder S) + private module Base = ≲-Reasoning (⊆-preorder S) open Base public hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨; step-≲; step-∼) diff --git a/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda index 3ea10495e6..17a0c25c77 100644 --- a/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda @@ -30,8 +30,8 @@ open import Relation.Binary.Definitions as B open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; _≢_; refl; sym; subst; subst₂) -import Data.List.Properties as Listₚ -import Data.List.Relation.Binary.Prefix.Heterogeneous.Properties as Prefixₚ +import Data.List.Properties as List +import Data.List.Relation.Binary.Prefix.Heterogeneous.Properties as Prefix ------------------------------------------------------------------------ -- Suffix and Prefix are linked via reverse @@ -43,15 +43,15 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where fromPrefix {as} {bs} p with Prefix.toView p ... | Prefix._++_ {cs} rs ds = subst (Suffix R (reverse as)) - (sym (Listₚ.reverse-++ cs ds)) + (sym (List.reverse-++ cs ds)) (Suffix.fromView (reverse ds Suffix.++ Pw.reverse⁺ rs)) fromPrefix-rev : ∀ {as bs} → Prefix R (reverse as) (reverse bs) → Suffix R as bs fromPrefix-rev pre = subst₂ (Suffix R) - (Listₚ.reverse-involutive _) - (Listₚ.reverse-involutive _) + (List.reverse-involutive _) + (List.reverse-involutive _) (fromPrefix pre) toPrefix-rev : ∀ {as bs} → Suffix R as bs → @@ -59,15 +59,15 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where toPrefix-rev {as} {bs} s with Suffix.toView s ... | Suffix._++_ cs {ds} rs = subst (Prefix R (reverse as)) - (sym (Listₚ.reverse-++ cs ds)) + (sym (List.reverse-++ cs ds)) (Prefix.fromView (Pw.reverse⁺ rs Prefix.++ reverse cs)) toPrefix : ∀ {as bs} → Suffix R (reverse as) (reverse bs) → Prefix R as bs toPrefix suf = subst₂ (Prefix R) - (Listₚ.reverse-involutive _) - (Listₚ.reverse-involutive _) + (List.reverse-involutive _) + (List.reverse-involutive _) (toPrefix-rev suf) ------------------------------------------------------------------------ @@ -139,7 +139,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where ds_; z0 {1+[2 _ ]} _ = 0 0ᵇ → x < x + y x0 = begin-strict x ≡⟨ sym (fromℕ-toℕ x) ⟩ - fromℕ (toℕ x) <⟨ fromℕ-mono-< (ℕₚ.m0)) ⟩ + fromℕ (toℕ x) <⟨ fromℕ-mono-< (ℕ.m0)) ⟩ fromℕ (toℕ x ℕ.+ toℕ y) ≡⟨ fromℕ-homo-+ (toℕ x) (toℕ y) ⟩ fromℕ (toℕ x) + fromℕ (toℕ y) ≡⟨ cong₂ _+_ (fromℕ-toℕ x) (fromℕ-toℕ y) ⟩ x + y ∎ @@ -910,12 +910,12 @@ x≢0⇒x+y≢0 {zero} _ 0≢0 = contradiction refl 0≢0 private 2*ₙ2*ₙ = (2 ℕ.*_) ∘ (2 ℕ.*_) toℕ-homo-* : ∀ x y → toℕ (x * y) ≡ toℕ x ℕ.* toℕ y -toℕ-homo-* x y = aux x y (size x ℕ.+ size y) ℕₚ.≤-refl +toℕ-homo-* x y = aux x y (size x ℕ.+ size y) ℕ.≤-refl where aux : (x y : ℕᵇ) → (cnt : ℕ) → (size x ℕ.+ size y ℕ.≤ cnt) → toℕ (x * y) ≡ toℕ x ℕ.* toℕ y aux zero _ _ _ = refl - aux 2[1+ x ] zero _ _ = sym (ℕₚ.*-zeroʳ (toℕ x ℕ.+ (ℕ.suc (toℕ x ℕ.+ 0)))) - aux 1+[2 x ] zero _ _ = sym (ℕₚ.*-zeroʳ (toℕ x ℕ.+ (toℕ x ℕ.+ 0))) + aux 2[1+ x ] zero _ _ = sym (ℕ.*-zeroʳ (toℕ x ℕ.+ (ℕ.suc (toℕ x ℕ.+ 0)))) + aux 1+[2 x ] zero _ _ = sym (ℕ.*-zeroʳ (toℕ x ℕ.+ (toℕ x ℕ.+ 0))) aux 2[1+ x ] 2[1+ y ] (ℕ.suc cnt) (s≤s |x|+1+|y|≤cnt) = begin toℕ (2[1+ x ] * 2[1+ y ]) ≡⟨⟩ toℕ (double 2[1+ (x + (y + xy)) ]) ≡⟨ toℕ-double 2[1+ (x + (y + xy)) ] ⟩ @@ -932,7 +932,7 @@ toℕ-homo-* x y = aux x y (size x ℕ.+ size y) ℕₚ.≤-refl where open ≡-Reasoning; m = toℕ x; n = toℕ y; xy = x * y - |x|+|y|≤cnt = ℕₚ.≤-trans (ℕₚ.+-monoʳ-≤ (size x) (ℕₚ.n≤1+n (size y))) |x|+1+|y|≤cnt + |x|+|y|≤cnt = ℕ.≤-trans (ℕ.+-monoʳ-≤ (size x) (ℕ.n≤1+n (size y))) |x|+1+|y|≤cnt aux 2[1+ x ] 1+[2 y ] (ℕ.suc cnt) (s≤s |x|+1+|y|≤cnt) = begin toℕ (2[1+ x ] * 1+[2 y ]) ≡⟨⟩ @@ -980,13 +980,13 @@ toℕ-homo-* x y = aux x y (size x ℕ.+ size y) ℕₚ.≤-refl ℕ.suc (2 ℕ.* (toℕ (x + y * 1+2x))) ≡⟨ cong (ℕ.suc ∘ (2 ℕ.*_)) (toℕ-homo-+ x (y * 1+2x)) ⟩ ℕ.suc (2 ℕ.* (m ℕ.+ (toℕ (y * 1+2x)))) ≡⟨ cong (ℕ.suc ∘ (2 ℕ.*_) ∘ (m ℕ.+_)) (aux y 1+2x cnt |y|+1+|x|≤cnt) ⟩ - ℕ.suc (2 ℕ.* (m ℕ.+ (n ℕ.* [1+2x]′))) ≡⟨ cong ℕ.suc $ ℕₚ.*-distribˡ-+ 2 m (n ℕ.* [1+2x]′) ⟩ - ℕ.suc (2m ℕ.+ (2 ℕ.* (n ℕ.* [1+2x]′))) ≡⟨ cong (ℕ.suc ∘ (2m ℕ.+_)) (sym (ℕₚ.*-assoc 2 n _)) ⟩ + ℕ.suc (2 ℕ.* (m ℕ.+ (n ℕ.* [1+2x]′))) ≡⟨ cong ℕ.suc $ ℕ.*-distribˡ-+ 2 m (n ℕ.* [1+2x]′) ⟩ + ℕ.suc (2m ℕ.+ (2 ℕ.* (n ℕ.* [1+2x]′))) ≡⟨ cong (ℕ.suc ∘ (2m ℕ.+_)) (sym (ℕ.*-assoc 2 n _)) ⟩ (ℕ.suc 2m) ℕ.+ 2n ℕ.* [1+2x]′ ≡⟨⟩ [1+2x]′ ℕ.+ 2n ℕ.* [1+2x]′ ≡⟨ cong (ℕ._+ (2n ℕ.* [1+2x]′)) $ - sym (ℕₚ.*-identityˡ [1+2x]′) ⟩ - 1 ℕ.* [1+2x]′ ℕ.+ 2n ℕ.* [1+2x]′ ≡⟨ sym (ℕₚ.*-distribʳ-+ [1+2x]′ 1 2n) ⟩ - (ℕ.suc 2n) ℕ.* [1+2x]′ ≡⟨ ℕₚ.*-comm (ℕ.suc 2n) [1+2x]′ ⟩ + sym (ℕ.*-identityˡ [1+2x]′) ⟩ + 1 ℕ.* [1+2x]′ ℕ.+ 2n ℕ.* [1+2x]′ ≡⟨ sym (ℕ.*-distribʳ-+ [1+2x]′ 1 2n) ⟩ + (ℕ.suc 2n) ℕ.* [1+2x]′ ≡⟨ ℕ.*-comm (ℕ.suc 2n) [1+2x]′ ⟩ toℕ 1+[2 x ] ℕ.* toℕ 1+[2 y ] ∎ where open ≡-Reasoning @@ -999,19 +999,19 @@ toℕ-homo-* x y = aux x y (size x ℕ.+ size y) ℕₚ.≤-refl |y|+1+|x|≤cnt = subst (ℕ._≤ cnt) eq |x|+1+|y|≤cnt -toℕ-isMagmaHomomorphism-* : IsMagmaHomomorphism *-rawMagma ℕᵇ.*-rawMagma toℕ +toℕ-isMagmaHomomorphism-* : IsMagmaHomomorphism *-rawMagma ℕ.*-rawMagma toℕ toℕ-isMagmaHomomorphism-* = record { isRelHomomorphism = toℕ-isRelHomomorphism ; homo = toℕ-homo-* } -toℕ-isMonoidHomomorphism-* : IsMonoidHomomorphism *-1-rawMonoid ℕᵇ.*-1-rawMonoid toℕ +toℕ-isMonoidHomomorphism-* : IsMonoidHomomorphism *-1-rawMonoid ℕ.*-1-rawMonoid toℕ toℕ-isMonoidHomomorphism-* = record { isMagmaHomomorphism = toℕ-isMagmaHomomorphism-* ; ε-homo = refl } -toℕ-isMonoidMonomorphism-* : IsMonoidMonomorphism *-1-rawMonoid ℕᵇ.*-1-rawMonoid toℕ +toℕ-isMonoidMonomorphism-* : IsMonoidMonomorphism *-1-rawMonoid ℕ.*-1-rawMonoid toℕ toℕ-isMonoidMonomorphism-* = record { isMonoidHomomorphism = toℕ-isMonoidHomomorphism-* ; injective = toℕ-injective @@ -1038,13 +1038,13 @@ private -- by `toℕ`/`fromℕ`. *-assoc : Associative _*_ -*-assoc = *-Monomorphism.assoc ℕₚ.*-isMagma ℕₚ.*-assoc +*-assoc = *-Monomorphism.assoc ℕ.*-isMagma ℕ.*-assoc *-comm : Commutative _*_ -*-comm = *-Monomorphism.comm ℕₚ.*-isMagma ℕₚ.*-comm +*-comm = *-Monomorphism.comm ℕ.*-isMagma ℕ.*-comm *-identityˡ : LeftIdentity 1ᵇ _*_ -*-identityˡ = *-Monomorphism.identityˡ ℕₚ.*-isMagma ℕₚ.*-identityˡ +*-identityˡ = *-Monomorphism.identityˡ ℕ.*-isMagma ℕ.*-identityˡ *-identityʳ : RightIdentity 1ᵇ _*_ *-identityʳ x = trans (*-comm x 1ᵇ) (*-identityˡ x) @@ -1068,7 +1068,7 @@ private a * (b + c) ≡⟨ sym (fromℕ-toℕ (a * (b + c))) ⟩ fromℕ (toℕ (a * (b + c))) ≡⟨ cong fromℕ (toℕ-homo-* a (b + c)) ⟩ fromℕ (k ℕ.* (toℕ (b + c))) ≡⟨ cong (fromℕ ∘ (k ℕ.*_)) (toℕ-homo-+ b c) ⟩ - fromℕ (k ℕ.* (m ℕ.+ n)) ≡⟨ cong fromℕ (ℕₚ.*-distribˡ-+ k m n) ⟩ + fromℕ (k ℕ.* (m ℕ.+ n)) ≡⟨ cong fromℕ (ℕ.*-distribˡ-+ k m n) ⟩ fromℕ (k ℕ.* m ℕ.+ k ℕ.* n) ≡⟨ cong fromℕ $ sym $ cong₂ ℕ._+_ (toℕ-homo-* a b) (toℕ-homo-* a c) ⟩ fromℕ (toℕ (a * b) ℕ.+ toℕ (a * c)) ≡⟨ cong fromℕ (sym (toℕ-homo-+ (a * b) (a * c))) ⟩ @@ -1089,13 +1089,13 @@ private *-isMagma = isMagma _*_ *-isSemigroup : IsSemigroup _*_ -*-isSemigroup = *-Monomorphism.isSemigroup ℕₚ.*-isSemigroup +*-isSemigroup = *-Monomorphism.isSemigroup ℕ.*-isSemigroup *-1-isMonoid : IsMonoid _*_ 1ᵇ -*-1-isMonoid = *-Monomorphism.isMonoid ℕₚ.*-1-isMonoid +*-1-isMonoid = *-Monomorphism.isMonoid ℕ.*-1-isMonoid *-1-isCommutativeMonoid : IsCommutativeMonoid _*_ 1ᵇ -*-1-isCommutativeMonoid = *-Monomorphism.isCommutativeMonoid ℕₚ.*-1-isCommutativeMonoid +*-1-isCommutativeMonoid = *-Monomorphism.isCommutativeMonoid ℕ.*-1-isCommutativeMonoid +-*-isSemiringWithoutAnnihilatingZero : IsSemiringWithoutAnnihilatingZero _+_ _*_ zero 1ᵇ +-*-isSemiringWithoutAnnihilatingZero = record @@ -1157,10 +1157,10 @@ private *-mono-≤ : _*_ Preserves₂ _≤_ ⟶ _≤_ ⟶ _≤_ *-mono-≤ {x} {u} {y} {v} x≤u y≤v = toℕ-cancel-≤ (begin toℕ (x * y) ≡⟨ toℕ-homo-* x y ⟩ - toℕ x ℕ.* toℕ y ≤⟨ ℕₚ.*-mono-≤ (toℕ-mono-≤ x≤u) (toℕ-mono-≤ y≤v) ⟩ + toℕ x ℕ.* toℕ y ≤⟨ ℕ.*-mono-≤ (toℕ-mono-≤ x≤u) (toℕ-mono-≤ y≤v) ⟩ toℕ u ℕ.* toℕ v ≡⟨ sym (toℕ-homo-* u v) ⟩ toℕ (u * v) ∎) - where open ℕₚ.≤-Reasoning + where open ℕ.≤-Reasoning *-monoʳ-≤ : ∀ x → (x *_) Preserves _≤_ ⟶ _≤_ *-monoʳ-≤ x y≤y′ = *-mono-≤ (≤-refl {x}) y≤y′ @@ -1171,10 +1171,10 @@ private *-mono-< : _*_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_ *-mono-< {x} {u} {y} {v} x {x} {y} x>y with x ≤? y ... | yes x≤y = contradiction x>y (≤⇒≯ x≤y) x≤y⇒x∸y≡0 : x ≤ y → x ∸ y ≡ 0ᵇ -x≤y⇒x∸y≡0 {x} {y} = toℕ-injective ∘ trans (toℕ-homo-∸ x y) ∘ ℕₚ.m≤n⇒m∸n≡0 ∘ toℕ-mono-≤ +x≤y⇒x∸y≡0 {x} {y} = toℕ-injective ∘ trans (toℕ-homo-∸ x y) ∘ ℕ.m≤n⇒m∸n≡0 ∘ toℕ-mono-≤ x∸y≡0⇒x≤y : x ∸ y ≡ 0ᵇ → x ≤ y -x∸y≡0⇒x≤y {x} {y} = toℕ-cancel-≤ ∘ ℕₚ.m∸n≡0⇒m≤n ∘ trans (sym (toℕ-homo-∸ x y)) ∘ cong toℕ +x∸y≡0⇒x≤y {x} {y} = toℕ-cancel-≤ ∘ ℕ.m∸n≡0⇒m≤n ∘ trans (sym (toℕ-homo-∸ x y)) ∘ cong toℕ x0 : x < y → y ∸ x > 0ᵇ -x0 {x} {y} = toℕ-cancel-< ∘ subst (ℕ._> 0) (sym (toℕ-homo-∸ y x)) ∘ ℕₚ.m0 {x} {y} = toℕ-cancel-< ∘ subst (ℕ._> 0) (sym (toℕ-homo-∸ y x)) ∘ ℕ.m0 {x} {y} = toℕ-cancel-< ∘ subst (ℕ._> 0) (sym (toℕ-homo-∸ [x∸y]+y≡x {x} {y} x≥y = toℕ-injective (begin toℕ (x ∸ y + y) ≡⟨ toℕ-homo-+ (x ∸ y) y ⟩ toℕ (x ∸ y) ℕ.+ toℕ y ≡⟨ cong (ℕ._+ toℕ y) (toℕ-homo-∸ x y) ⟩ - (toℕ x ℕ.∸ toℕ y) ℕ.+ toℕ y ≡⟨ ℕₚ.m∸n+n≡m (toℕ-mono-≤ x≥y) ⟩ + (toℕ x ℕ.∸ toℕ y) ℕ.+ toℕ y ≡⟨ ℕ.m∸n+n≡m (toℕ-mono-≤ x≥y) ⟩ toℕ x ∎) where open ≡-Reasoning @@ -162,7 +162,7 @@ x+[y∸x]≡y {x} {y} x≤y = begin-equality ∸-+-assoc x y z = toℕ-injective $ begin toℕ ((x ∸ y) ∸ z) ≡⟨ toℕ-homo-∸ (x ∸ y) z ⟩ toℕ (x ∸ y) ℕ.∸ n ≡⟨ cong (ℕ._∸ n) (toℕ-homo-∸ x y) ⟩ - (k ℕ.∸ m) ℕ.∸ n ≡⟨ ℕₚ.∸-+-assoc k m n ⟩ + (k ℕ.∸ m) ℕ.∸ n ≡⟨ ℕ.∸-+-assoc k m n ⟩ k ℕ.∸ (m ℕ.+ n) ≡⟨ cong (k ℕ.∸_) (sym (toℕ-homo-+ y z)) ⟩ k ℕ.∸ (toℕ (y + z)) ≡⟨ sym (toℕ-homo-∸ x (y + z)) ⟩ toℕ (x ∸ (y + z)) ∎ @@ -172,7 +172,7 @@ x+[y∸x]≡y {x} {y} x≤y = begin-equality +-∸-assoc x {y} {z} z≤y = toℕ-injective $ begin toℕ ((x + y) ∸ z) ≡⟨ toℕ-homo-∸ (x + y) z ⟩ (toℕ (x + y)) ℕ.∸ n ≡⟨ cong (ℕ._∸ n) (toℕ-homo-+ x y) ⟩ - (k ℕ.+ m) ℕ.∸ n ≡⟨ ℕₚ.+-∸-assoc k n≤m ⟩ + (k ℕ.+ m) ℕ.∸ n ≡⟨ ℕ.+-∸-assoc k n≤m ⟩ k ℕ.+ (m ℕ.∸ n) ≡⟨ cong (k ℕ.+_) (sym (toℕ-homo-∸ y z)) ⟩ k ℕ.+ toℕ (y ∸ z) ≡⟨ sym (toℕ-homo-+ x (y ∸ z)) ⟩ toℕ (x + (y ∸ z)) ∎ @@ -196,10 +196,10 @@ suc[x]∸suc[y] x y = begin-equality ∸-mono-≤ : _∸_ Preserves₂ _≤_ ⟶ _≥_ ⟶ _≤_ ∸-mono-≤ {x} {y} {u} {v} x≤y v≤u = toℕ-cancel-≤ (begin toℕ (x ∸ u) ≡⟨ toℕ-homo-∸ x u ⟩ - toℕ x ℕ.∸ toℕ u ≤⟨ ℕₚ.∸-mono (toℕ-mono-≤ x≤y) (toℕ-mono-≤ v≤u) ⟩ + toℕ x ℕ.∸ toℕ u ≤⟨ ℕ.∸-mono (toℕ-mono-≤ x≤y) (toℕ-mono-≤ v≤u) ⟩ toℕ y ℕ.∸ toℕ v ≡⟨ sym (toℕ-homo-∸ y v) ⟩ toℕ (y ∸ v) ∎) - where open ℕₚ.≤-Reasoning + where open ℕ.≤-Reasoning ∸-monoˡ-≤ : (x : ℕᵇ) → (_∸ x) Preserves _≤_ ⟶ _≤_ ∸-monoˡ-≤ x y≤z = ∸-mono-≤ y≤z (≤-refl {x}) @@ -230,7 +230,7 @@ x∸yₛ _×-function_ : Func A B → Func C D → Func (A ×ₛ C) (B ×ₛ D) f ×-function g = record - { to = Prod.map (to f) (to g) - ; cong = Prod.map (cong f) (cong g) + { to = Product.map (to f) (to g) + ; cong = Product.map (cong f) (cong g) } where open Func infixr 2 _×-equivalence_ _×-injection_ _×-left-inverse_ @@ -53,54 +53,54 @@ infixr 2 _×-equivalence_ _×-injection_ _×-left-inverse_ _×-equivalence_ : Equivalence A B → Equivalence C D → Equivalence (A ×ₛ C) (B ×ₛ D) _×-equivalence_ f g = record - { to = Prod.map (to f) (to g) - ; from = Prod.map (from f) (from g) - ; to-cong = Prod.map (to-cong f) (to-cong g) - ; from-cong = Prod.map (from-cong f) (from-cong g) + { to = Product.map (to f) (to g) + ; from = Product.map (from f) (from g) + ; to-cong = Product.map (to-cong f) (to-cong g) + ; from-cong = Product.map (from-cong f) (from-cong g) } where open Equivalence _×-injection_ : Injection A B → Injection C D → Injection (A ×ₛ C) (B ×ₛ D) f ×-injection g = record - { to = Prod.map (to f) (to g) - ; cong = Prod.map (cong f) (cong g) - ; injective = Prod.map (injective f) (injective g) + { to = Product.map (to f) (to g) + ; cong = Product.map (cong f) (cong g) + ; injective = Product.map (injective f) (injective g) } where open Injection _×-surjection_ : Surjection A B → Surjection C D → Surjection (A ×ₛ C) (B ×ₛ D) f ×-surjection g = record - { to = Prod.map (to f) (to g) - ; cong = Prod.map (cong f) (cong g) - ; surjective = λ y → Prod.zip _,_ (λ ff gg x₂ → (ff (proj₁ x₂)) , (gg (proj₂ x₂))) (surjective f (proj₁ y)) (surjective g (proj₂ y)) + { to = Product.map (to f) (to g) + ; cong = Product.map (cong f) (cong g) + ; surjective = λ y → Product.zip _,_ (λ ff gg x₂ → (ff (proj₁ x₂)) , (gg (proj₂ x₂))) (surjective f (proj₁ y)) (surjective g (proj₂ y)) } where open Surjection _×-bijection_ : Bijection A B → Bijection C D → Bijection (A ×ₛ C) (B ×ₛ D) f ×-bijection g = record - { to = Prod.map (to f) (to g) - ; cong = Prod.map (cong f) (cong g) - ; bijective = Prod.map (injective f) (injective g) , - λ { (y₀ , y₁) → Prod.zip _,_ (λ {ff gg (x₀ , x₁) → ff x₀ , gg x₁}) (surjective f y₀) (surjective g y₁)} + { to = Product.map (to f) (to g) + ; cong = Product.map (cong f) (cong g) + ; bijective = Product.map (injective f) (injective g) , + λ { (y₀ , y₁) → Product.zip _,_ (λ {ff gg (x₀ , x₁) → ff x₀ , gg x₁}) (surjective f y₀) (surjective g y₁)} } where open Bijection _×-leftInverse_ : LeftInverse A B → LeftInverse C D → LeftInverse (A ×ₛ C) (B ×ₛ D) f ×-leftInverse g = record - { to = Prod.map (to f) (to g) - ; from = Prod.map (from f) (from g) - ; to-cong = Prod.map (to-cong f) (to-cong g) - ; from-cong = Prod.map (from-cong f) (from-cong g) + { to = Product.map (to f) (to g) + ; from = Product.map (from f) (from g) + ; to-cong = Product.map (to-cong f) (to-cong g) + ; from-cong = Product.map (from-cong f) (from-cong g) ; inverseˡ = λ x → inverseˡ f (proj₁ x) , inverseˡ g (proj₂ x) } where open LeftInverse _×-rightInverse_ : RightInverse A B → RightInverse C D → RightInverse (A ×ₛ C) (B ×ₛ D) f ×-rightInverse g = record - { to = Prod.map (to f) (to g) - ; from = Prod.map (from f) (from g) - ; to-cong = Prod.map (to-cong f) (to-cong g) - ; from-cong = Prod.map (from-cong f) (from-cong g) + { to = Product.map (to f) (to g) + ; from = Product.map (from f) (from g) + ; to-cong = Product.map (to-cong f) (to-cong g) + ; from-cong = Product.map (from-cong f) (from-cong g) ; inverseʳ = λ x → inverseʳ f (proj₁ x) , inverseʳ g (proj₂ x) } where open RightInverse @@ -109,10 +109,10 @@ infixr 2 _×-surjection_ _×-inverse_ _×-inverse_ : Inverse A B → Inverse C D → Inverse (A ×ₛ C) (B ×ₛ D) f ×-inverse g = record - { to = Prod.map (to f) (to g) - ; from = Prod.map (from f) (from g) - ; to-cong = Prod.map (to-cong f) (to-cong g) - ; from-cong = Prod.map (from-cong f) (from-cong g) + { to = Product.map (to f) (to g) + ; from = Product.map (from f) (from g) + ; to-cong = Product.map (to-cong f) (to-cong g) + ; from-cong = Product.map (from-cong f) (from-cong g) ; inverse = (λ x → inverseˡ f (proj₁ x) , inverseˡ g (proj₂ x)) , (λ x → inverseʳ f (proj₁ x) , inverseʳ g (proj₂ x)) } where open Inverse diff --git a/src/Data/Product/Relation/Binary/Pointwise/Dependent.agda b/src/Data/Product/Relation/Binary/Pointwise/Dependent.agda index fd22ff6cce..ac0f48df7b 100644 --- a/src/Data/Product/Relation/Binary/Pointwise/Dependent.agda +++ b/src/Data/Product/Relation/Binary/Pointwise/Dependent.agda @@ -8,7 +8,7 @@ module Data.Product.Relation.Binary.Pointwise.Dependent where -open import Data.Product.Base as Prod +open import Data.Product.Base as Product open import Level open import Function.Base open import Relation.Binary.Core using (Rel; REL; _⇒_) diff --git a/src/Data/Product/Relation/Binary/Pointwise/NonDependent.agda b/src/Data/Product/Relation/Binary/Pointwise/NonDependent.agda index efa64a8605..95a8978cf5 100644 --- a/src/Data/Product/Relation/Binary/Pointwise/NonDependent.agda +++ b/src/Data/Product/Relation/Binary/Pointwise/NonDependent.agda @@ -8,7 +8,7 @@ module Data.Product.Relation.Binary.Pointwise.NonDependent where -open import Data.Product.Base as Prod +open import Data.Product.Base as Product open import Data.Product.Properties using (≡-dec) open import Data.Sum.Base open import Data.Unit.Base using (⊤) @@ -35,7 +35,7 @@ Pointwise R S = (R on proj₁) -×- (S on proj₂) -- Pointwise preserves many relational properties ×-reflexive : ≈₁ ⇒ R → ≈₂ ⇒ S → Pointwise ≈₁ ≈₂ ⇒ Pointwise R S -×-reflexive refl₁ refl₂ = Prod.map refl₁ refl₂ +×-reflexive refl₁ refl₂ = Product.map refl₁ refl₂ ×-refl : Reflexive R → Reflexive S → Reflexive (Pointwise R S) ×-refl refl₁ refl₂ = refl₁ , refl₂ @@ -49,14 +49,14 @@ Pointwise R S = (R on proj₁) -×- (S on proj₂) ×-irreflexive₂ ir x≈y x-nonZero : ∀ {p} → p > 0ℚ → NonZero p >-nonZero {p@(mkℚ _ _ _)} (*<* p-nonZero {toℚᵘ p} (ℚᵘ.*<* p; uncurry) open import Data.Sum.Base using ([_,_]′) open import Data.Sum.Properties using ([,]-map) @@ -689,7 +689,7 @@ map-<,>-zip f g [] = refl map-<,>-zip f g (x ∷ xs) = cong (_ ∷_) (map-<,>-zip f g xs) map-zip : ∀ (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 (Product.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) = cong (_ ∷_) (map-zip f g xs ys) @@ -705,10 +705,10 @@ lookup-unzip (suc i) ((x , y) ∷ xys) = lookup-unzip i xys map-unzip : ∀ (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) + in (map f xs , map g ys) ≡ unzip (map (Product.map f g) xys) map-unzip f g [] = refl map-unzip f g ((x , y) ∷ xys) = - cong (Prod.map (f x ∷_) (g y ∷_)) (map-unzip f g xys) + cong (Product.map (f x ∷_) (g y ∷_)) (map-unzip f g xys) -- Products of vectors are isomorphic to vectors of products. @@ -716,7 +716,7 @@ unzip∘zip : ∀ (xs : Vec A n) (ys : Vec B n) → unzip (zip xs ys) ≡ (xs , ys) unzip∘zip [] [] = refl unzip∘zip (x ∷ xs) (y ∷ ys) = - cong (Prod.map (x ∷_) (y ∷_)) (unzip∘zip xs ys) + cong (Product.map (x ∷_) (y ∷_)) (unzip∘zip xs ys) zip∘unzip : ∀ (xys : Vec (A × B) n) → uncurry zip (unzip xys) ≡ xys zip∘unzip [] = refl @@ -869,7 +869,7 @@ unfold-∷ʳ eq x (y ∷ xs) = cong (y ∷_) (unfold-∷ʳ (cong pred eq) x xs) ∷ʳ-injective : ∀ (xs ys : Vec A n) → 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₁ (cong (x ∷_)) (∷ʳ-injective xs ys eq′) +... | refl , eq′ = Product.map₁ (cong (x ∷_)) (∷ʳ-injective xs ys eq′) ∷ʳ-injectiveˡ : ∀ (xs ys : Vec A n) → xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys ∷ʳ-injectiveˡ xs ys eq = proj₁ (∷ʳ-injective xs ys eq) @@ -1294,29 +1294,29 @@ cast-fromList : ∀ {xs ys : List A} (eq : xs ≡ ys) → cast (cong List.length eq) (fromList xs) ≡ fromList ys cast-fromList {xs = List.[]} {ys = List.[]} eq = refl cast-fromList {xs = x List.∷ xs} {ys = y List.∷ ys} eq = - let x≡y , xs≡ys = Listₚ.∷-injective eq in begin + let x≡y , xs≡ys = List.∷-injective eq in begin x ∷ cast (cong (pred ∘ List.length) eq) (fromList xs) ≡⟨ cong (_ ∷_) (cast-fromList xs≡ys) ⟩ x ∷ fromList ys ≡⟨ cong (_∷ _) x≡y ⟩ y ∷ fromList ys ∎ where open ≡-Reasoning fromList-map : ∀ (f : A → B) (xs : List A) → - cast (Listₚ.length-map f xs) (fromList (List.map f xs)) ≡ map f (fromList xs) + cast (List.length-map f xs) (fromList (List.map f xs)) ≡ map f (fromList xs) fromList-map f List.[] = refl fromList-map f (x List.∷ xs) = cong (f x ∷_) (fromList-map f xs) fromList-++ : ∀ (xs : List A) {ys : List A} → - cast (Listₚ.length-++ xs) (fromList (xs List.++ ys)) ≡ fromList xs ++ fromList ys + cast (List.length-++ xs) (fromList (xs List.++ ys)) ≡ fromList xs ++ fromList ys fromList-++ List.[] {ys} = cast-is-id refl (fromList ys) fromList-++ (x List.∷ xs) {ys} = cong (x ∷_) (fromList-++ xs) -fromList-reverse : (xs : List A) → cast (Listₚ.length-reverse xs) (fromList (List.reverse xs)) ≡ reverse (fromList xs) +fromList-reverse : (xs : List A) → cast (List.length-reverse xs) (fromList (List.reverse xs)) ≡ reverse (fromList xs) fromList-reverse List.[] = refl fromList-reverse (x List.∷ xs) = begin - fromList (List.reverse (x List.∷ xs)) ≈⟨ cast-fromList (Listₚ.ʳ++-defn xs) ⟩ + fromList (List.reverse (x List.∷ xs)) ≈⟨ cast-fromList (List.ʳ++-defn xs) ⟩ fromList (List.reverse xs List.++ List.[ x ]) ≈⟨ fromList-++ (List.reverse xs) ⟩ fromList (List.reverse xs) ++ [ x ] ≈⟨ unfold-∷ʳ (+-comm 1 _) x (fromList (List.reverse xs)) ⟨ - fromList (List.reverse xs) ∷ʳ x ≈⟨ ≈-cong (_∷ʳ x) (cast-∷ʳ (cong suc (Listₚ.length-reverse xs)) _ _) + fromList (List.reverse xs) ∷ʳ x ≈⟨ ≈-cong (_∷ʳ x) (cast-∷ʳ (cong suc (List.length-reverse xs)) _ _) (fromList-reverse xs) ⟩ reverse (fromList xs) ∷ʳ x ≂⟨ reverse-∷ x (fromList xs) ⟨ reverse (x ∷ fromList xs) ≈⟨⟩ diff --git a/src/Data/Vec/Recursive.agda b/src/Data/Vec/Recursive.agda index ea7d04285f..0eb04b264b 100644 --- a/src/Data/Vec/Recursive.agda +++ b/src/Data/Vec/Recursive.agda @@ -21,7 +21,7 @@ open import Data.Nat.Properties using (+-comm; *-comm) open import Data.Empty.Polymorphic open import Data.Fin.Base as Fin using (Fin; zero; suc) open import Data.Fin.Properties using (1↔⊤; *↔×) -open import Data.Product.Base as Prod using (_×_; _,_; proj₁; proj₂) +open import Data.Product.Base as Product using (_×_; _,_; proj₁; proj₂) open import Data.Product.Algebra using (×-cong) open import Data.Sum.Base as Sum using (_⊎_) open import Data.Unit.Base using (tt) @@ -146,7 +146,7 @@ zipWith f ((suc n@(suc _))) (a , as) (b , bs) = f a b , zipWith f 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 (suc n@(suc _)) (a , as) = Prod.zip _,_ _,_ (f a) (unzipWith f n as) +unzipWith f (suc n@(suc _)) (a , as) = Product.zip _,_ _,_ (f a) (unzipWith f n as) zip : ∀ n → A ^ n → B ^ n → (A × B) ^ n zip = zipWith _,_ diff --git a/src/Data/Vec/Relation/Unary/All.agda b/src/Data/Vec/Relation/Unary/All.agda index e8d24fbd65..4cc1805c25 100644 --- a/src/Data/Vec/Relation/Unary/All.agda +++ b/src/Data/Vec/Relation/Unary/All.agda @@ -9,7 +9,7 @@ module Data.Vec.Relation.Unary.All where open import Data.Nat.Base using (ℕ; zero; suc; NonZero) -open import Data.Product.Base as Prod using (_×_; _,_; uncurry; <_,_>) +open import Data.Product.Base as Product using (_×_; _,_; uncurry; <_,_>) open import Data.Sum.Base as Sum using (inj₁; inj₂) open import Data.Vec.Base as Vec using (Vec; []; _∷_) open import Data.Vec.Relation.Unary.Any as Any using (Any; here; there) @@ -71,7 +71,7 @@ zip (px ∷ pxs , qx ∷ qxs) = (px , qx) ∷ zip (pxs , qxs) unzip : All (P ∩ Q) {n} ⊆ All P ∩ All Q unzip [] = [] , [] -unzip (pqx ∷ pqxs) = Prod.zip _∷_ _∷_ pqx (unzip pqxs) +unzip (pqx ∷ pqxs) = Product.zip _∷_ _∷_ pqx (unzip pqxs) module _ {P : Pred A p} {Q : Pred B q} {R : Pred C r} where @@ -91,7 +91,7 @@ lookupAny (px ∷ pxs) (here qx) = px , qx lookupAny (px ∷ pxs) (there i) = lookupAny pxs i lookupWith : ∀[ P ⇒ Q ⇒ R ] → All P xs → (i : Any Q xs) → R (Any.lookup i) -lookupWith f pxs i = Prod.uncurry f (lookupAny pxs i) +lookupWith f pxs i = Product.uncurry f (lookupAny pxs i) lookup : All P xs → (∀ {x} → x ∈ₚ xs → P x) lookup pxs = lookupWith (λ { px P.refl → px }) pxs @@ -121,7 +121,7 @@ irrelevant irr (px₁ ∷ pxs₁) (px₂ ∷ pxs₂) = satisfiable : Satisfiable P → ∀ {n} → Satisfiable (All P {n}) satisfiable (x , p) {zero} = [] , [] -satisfiable (x , p) {suc n} = Prod.map (x ∷_) (p ∷_) (satisfiable (x , p)) +satisfiable (x , p) {suc n} = Product.map (x ∷_) (p ∷_) (satisfiable (x , p)) ------------------------------------------------------------------------ -- Generalised decidability procedure diff --git a/src/Data/Vec/Relation/Unary/All/Properties.agda b/src/Data/Vec/Relation/Unary/All/Properties.agda index 4b8951ff9d..3c2a126230 100644 --- a/src/Data/Vec/Relation/Unary/All/Properties.agda +++ b/src/Data/Vec/Relation/Unary/All/Properties.agda @@ -12,9 +12,8 @@ open import Data.Nat.Base using (ℕ; zero; suc; _+_) open import Data.Fin.Base using (Fin; zero; suc) open import Data.List.Base using ([]; _∷_) open import Data.List.Relation.Unary.All as List using ([]; _∷_) -open import Data.Product.Base as Prod using (_×_; _,_; uncurry; uncurry′) +open import Data.Product.Base as Product using (_×_; _,_; uncurry; uncurry′) open import Data.Vec.Base as Vec -import Data.Vec.Properties as Vecₚ open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) open import Level using (Level) open import Function.Base using (_∘_; id) @@ -82,7 +81,7 @@ gmap g = map⁺ ∘ All.map g ++⁻ : (xs : Vec A m) {ys : Vec A n} → All P (xs ++ ys) → All P xs × All P ys ++⁻ [] p = [] , p -++⁻ (x ∷ xs) (px ∷ pxs) = Prod.map₁ (px ∷_) (++⁻ _ pxs) +++⁻ (x ∷ xs) (px ∷ pxs) = Product.map₁ (px ∷_) (++⁻ _ pxs) ++⁺∘++⁻ : (xs : Vec A m) {ys : Vec A n} → (p : All P (xs ++ ys)) → diff --git a/src/Data/Vec/Relation/Unary/AllPairs/Properties.agda b/src/Data/Vec/Relation/Unary/AllPairs/Properties.agda index ffc0ec6c4c..b13962b750 100644 --- a/src/Data/Vec/Relation/Unary/AllPairs/Properties.agda +++ b/src/Data/Vec/Relation/Unary/AllPairs/Properties.agda @@ -9,9 +9,9 @@ module Data.Vec.Relation.Unary.AllPairs.Properties where open import Data.Vec -import Data.Vec.Properties as Vecₚ +import Data.Vec.Properties as Vec open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) -import Data.Vec.Relation.Unary.All.Properties as Allₚ +import Data.Vec.Relation.Unary.All.Properties as All open import Data.Vec.Relation.Unary.AllPairs as AllPairs using (AllPairs; []; _∷_) open import Data.Bool.Base using (true; false) open import Data.Fin.Base using (Fin) @@ -39,7 +39,7 @@ module _ {R : Rel A ℓ} {f : B → A} where map⁺ : ∀ {n xs} → AllPairs (λ x y → R (f x) (f y)) {n} xs → AllPairs R {n} (map f xs) map⁺ [] = [] - map⁺ (x∉xs ∷ xs!) = Allₚ.map⁺ x∉xs ∷ map⁺ xs! + map⁺ (x∉xs ∷ xs!) = All.map⁺ x∉xs ∷ map⁺ xs! ------------------------------------------------------------------------ -- ++ @@ -49,7 +49,7 @@ module _ {R : Rel A ℓ} where ++⁺ : ∀ {n m xs ys} → AllPairs R {n} xs → AllPairs R {m} ys → All (λ x → All (R x) ys) xs → AllPairs R (xs ++ ys) ++⁺ [] Rys _ = Rys - ++⁺ (px ∷ Rxs) Rys (Rxys ∷ Rxsys) = Allₚ.++⁺ px Rxys ∷ ++⁺ Rxs Rys Rxsys + ++⁺ (px ∷ Rxs) Rys (Rxys ∷ Rxsys) = All.++⁺ px Rxys ∷ ++⁺ Rxs Rys Rxsys ------------------------------------------------------------------------ -- concat @@ -61,7 +61,7 @@ module _ {R : Rel A ℓ} where AllPairs R (concat xss) concat⁺ [] [] = [] concat⁺ (pxs ∷ pxss) (Rxsxss ∷ Rxss) = ++⁺ pxs (concat⁺ pxss Rxss) - (All.map Allₚ.concat⁺ (Allₚ.All-swap Rxsxss)) + (All.map All.concat⁺ (All.All-swap Rxsxss)) ------------------------------------------------------------------------ -- take and drop @@ -70,7 +70,7 @@ module _ {R : Rel A ℓ} where take⁺ : ∀ {n} m {xs} → AllPairs R {m + n} xs → AllPairs R {m} (take m xs) take⁺ zero pxs = [] - take⁺ (suc m) {x ∷ xs} (px ∷ pxs) = Allₚ.take⁺ m px ∷ take⁺ m pxs + take⁺ (suc m) {x ∷ xs} (px ∷ pxs) = All.take⁺ m px ∷ take⁺ m pxs drop⁺ : ∀ {n} m {xs} → AllPairs R {m + n} xs → AllPairs R {n} (drop m xs) drop⁺ zero pxs = pxs @@ -85,5 +85,5 @@ module _ {R : Rel A ℓ} where AllPairs R (tabulate f) tabulate⁺ {zero} fᵢ~fⱼ = [] tabulate⁺ {suc n} fᵢ~fⱼ = - Allₚ.tabulate⁺ (λ j → fᵢ~fⱼ λ()) ∷ + All.tabulate⁺ (λ j → fᵢ~fⱼ λ()) ∷ tabulate⁺ (fᵢ~fⱼ ∘ (_∘ suc-injective)) diff --git a/src/Data/Vec/Relation/Unary/Any.agda b/src/Data/Vec/Relation/Unary/Any.agda index de901fe571..b4d857a69c 100644 --- a/src/Data/Vec/Relation/Unary/Any.agda +++ b/src/Data/Vec/Relation/Unary/Any.agda @@ -13,7 +13,7 @@ open import Data.Fin.Base using (Fin; zero; suc) open import Data.Nat.Base using (ℕ; zero; suc; NonZero) open import Data.Sum.Base using (_⊎_; inj₁; inj₂; [_,_]′) open import Data.Vec.Base as Vec using (Vec; []; [_]; _∷_) -open import Data.Product.Base as Prod using (∃; _,_) +open import Data.Product.Base as Product using (∃; _,_) open import Level using (Level; _⊔_) open import Relation.Nullary.Negation using (¬_; contradiction) open import Relation.Nullary.Decidable as Dec using (yes; no; _⊎-dec_) @@ -80,7 +80,7 @@ any? P? (x ∷ xs) = Dec.map′ fromSum toSum (P? x ⊎-dec 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)) +satisfiable (x , p) {suc n} = Product.map (x ∷_) there (satisfiable (x , p)) any = any? {-# WARNING_ON_USAGE any diff --git a/src/Data/Vec/Relation/Unary/Any/Properties.agda b/src/Data/Vec/Relation/Unary/Any/Properties.agda index 437e7933e7..d330ea5cef 100644 --- a/src/Data/Vec/Relation/Unary/Any/Properties.agda +++ b/src/Data/Vec/Relation/Unary/Any/Properties.agda @@ -15,7 +15,7 @@ open import Data.List.Base using ([]; _∷_) import Data.List.Relation.Unary.Any as List open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]′) open import Data.Sum.Function.Propositional using (_⊎-cong_) -open import Data.Product.Base as Prod using (∃; ∃₂; _×_; _,_; proj₁; proj₂) +open import Data.Product.Base as Product using (∃; ∃₂; _×_; _,_; proj₁; proj₂) open import Data.Vec.Base hiding (here; there) open import Data.Vec.Relation.Unary.Any as Any using (Any; here; there) open import Data.Vec.Membership.Propositional @@ -365,7 +365,7 @@ module _ {P : Pred A p} where tabulate⁻ : ∀ {n} {f : Fin n → A} → Any P (tabulate f) → ∃ λ i → P (f i) tabulate⁻ (here p) = zero , p - tabulate⁻ (there p) = Prod.map suc id (tabulate⁻ p) + tabulate⁻ (there p) = Product.map suc id (tabulate⁻ p) ------------------------------------------------------------------------ -- mapWith∈ @@ -384,7 +384,7 @@ module _ {P : Pred B p} where ∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs) mapWith∈⁻ (y ∷ xs) f (here p) = (y , here refl , p) mapWith∈⁻ (y ∷ xs) f (there p) = - Prod.map id (Prod.map there id) $ mapWith∈⁻ xs (f ∘ there) p + Product.map id (Product.map there id) $ mapWith∈⁻ xs (f ∘ there) p mapWith∈↔ : ∀ {n} {xs : Vec A n} {f : ∀ {x} → x ∈ xs → B} → (∃₂ λ x (x∈xs : x ∈ xs) → P (f x∈xs)) ↔ Any P (mapWith∈ xs f) diff --git a/src/Data/Vec/Relation/Unary/Linked.agda b/src/Data/Vec/Relation/Unary/Linked.agda index 368c22ddc1..c0f6b5fe89 100644 --- a/src/Data/Vec/Relation/Unary/Linked.agda +++ b/src/Data/Vec/Relation/Unary/Linked.agda @@ -11,7 +11,7 @@ module Data.Vec.Relation.Unary.Linked {a} {A : Set a} where open import Data.Nat.Base using (ℕ; zero; suc; _+_) open import Data.Vec.Base as Vec using (Vec; []; _∷_) open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) -open import Data.Product.Base as Prod using (_,_; _×_; uncurry; <_,_>) +open import Data.Product.Base as Product using (_,_; _×_; uncurry; <_,_>) open import Function.Base using (id; _∘_) open import Level using (Level; _⊔_) open import Relation.Binary.Core using (Rel; _⇒_) @@ -66,7 +66,7 @@ zipWith f (px ∷ pxs , qx ∷ qxs) = f (px , qx) ∷ zipWith f (pxs , qxs) unzipWith : R ⇒ P ∩ᵇ Q → Linked R {n} ⊆ Linked P ∩ᵘ Linked Q unzipWith f [] = [] , [] unzipWith f [-] = [-] , [-] -unzipWith f (rx ∷ rxs) = Prod.zip _∷_ _∷_ (f rx) (unzipWith f rxs) +unzipWith f (rx ∷ rxs) = Product.zip _∷_ _∷_ (f rx) (unzipWith f rxs) zip : Linked P {n} ∩ᵘ Linked Q ⊆ Linked (P ∩ᵇ Q) zip = zipWith id diff --git a/src/Data/Vec/Relation/Unary/Linked/Properties.agda b/src/Data/Vec/Relation/Unary/Linked/Properties.agda index 8ff5a4d2e3..e49cda76f6 100644 --- a/src/Data/Vec/Relation/Unary/Linked/Properties.agda +++ b/src/Data/Vec/Relation/Unary/Linked/Properties.agda @@ -10,7 +10,7 @@ module Data.Vec.Relation.Unary.Linked.Properties where open import Data.Vec.Base as Vec open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) -import Data.Vec.Relation.Unary.All.Properties as Allₚ +import Data.Vec.Relation.Unary.All.Properties as All open import Data.Vec.Relation.Unary.Linked as Linked using (Linked; []; [-]; _∷_) open import Data.Fin.Base using (zero; suc; _<_) @@ -43,7 +43,7 @@ module _ (trans : Transitive R) where lookup⁺ : ∀ {i j} {xs : Vec _ n} → Linked R xs → i < j → R (lookup xs i) (lookup xs j) - lookup⁺ {i = zero} {j = suc j} (rx ∷ rxs) i; proj₁; proj₂; ∃₂; ∃) open import Data.Product.Function.NonDependent.Propositional open import Data.Sum.Base as Sum @@ -47,12 +47,12 @@ private -- Σ is associative Σ-assoc : ∀ {A : Set a} {B : A → Set b} {C : (a : A) → B a → Set c} → Σ (Σ A B) (uncurry C) ↔ Σ A (λ a → Σ (B a) (C a)) -Σ-assoc = mk↔ₛ′ Prod.assocʳ Prod.assocˡ (λ _ → P.refl) (λ _ → P.refl) +Σ-assoc = mk↔ₛ′ Product.assocʳ Product.assocˡ (λ _ → P.refl) (λ _ → P.refl) -- × is commutative ×-comm : ∀ (A : Set a) (B : Set b) → (A × B) ↔ (B × A) -×-comm _ _ = mk↔ₛ′ Prod.swap Prod.swap (λ _ → P.refl) λ _ → P.refl +×-comm _ _ = mk↔ₛ′ Product.swap Product.swap (λ _ → P.refl) λ _ → P.refl -- × has ⊤ as its identity @@ -114,14 +114,14 @@ private ×-distribˡ-⊎ : ∀ ℓ → _DistributesOverˡ_ {ℓ = ℓ} _↔_ _×_ _⊎_ ×-distribˡ-⊎ ℓ _ _ _ = mk↔ₛ′ (uncurry λ x → [ inj₁ ∘′ (x ,_) , inj₂ ∘′ (x ,_) ]′) - [ Prod.map₂ inj₁ , Prod.map₂ inj₂ ]′ + [ Product.map₂ inj₁ , Product.map₂ inj₂ ]′ [ (λ _ → P.refl) , (λ _ → P.refl) ] (uncurry λ _ → [ (λ _ → P.refl) , (λ _ → P.refl) ]) ×-distribʳ-⊎ : ∀ ℓ → _DistributesOverʳ_ {ℓ = ℓ} _↔_ _×_ _⊎_ ×-distribʳ-⊎ ℓ _ _ _ = mk↔ₛ′ (uncurry [ curry inj₁ , curry inj₂ ]′) - [ Prod.map₁ inj₁ , Prod.map₁ inj₂ ]′ + [ Product.map₁ inj₁ , Product.map₁ inj₂ ]′ [ (λ _ → P.refl) , (λ _ → P.refl) ] (uncurry [ (λ _ _ → P.refl) , (λ _ _ → P.refl) ]) diff --git a/src/Induction/Lexicographic.agda b/src/Induction/Lexicographic.agda index 6acd96b7f1..d70b989acd 100644 --- a/src/Induction/Lexicographic.agda +++ b/src/Induction/Lexicographic.agda @@ -68,13 +68,13 @@ RecA ⊗ RecB = Σ-Rec RecA (λ _ → RecB) private open import Data.Nat.Base - open import Data.Nat.Induction as N + open import Data.Nat.Induction as ℕ -- The Ackermann function à la Rózsa Péter. ackermann : ℕ → ℕ → ℕ ackermann m n = - build [ N.recBuilder ⊗ N.recBuilder ] + build [ ℕ.recBuilder ⊗ ℕ.recBuilder ] (λ _ → ℕ) (λ { (zero , n) _ → 1 + n ; (suc m , zero) (_ , ackm•) → ackm• 1 diff --git a/src/Reflection/AST/Definition.agda b/src/Reflection/AST/Definition.agda index b97fe47cae..3f06fef070 100644 --- a/src/Reflection/AST/Definition.agda +++ b/src/Reflection/AST/Definition.agda @@ -8,8 +8,8 @@ module Reflection.AST.Definition where -import Data.List.Properties as Listₚ using (≡-dec) -import Data.Nat.Properties as ℕₚ using (_≟_) +import Data.List.Properties as List using (≡-dec) +import Data.Nat.Properties as ℕ using (_≟_) open import Data.Product.Base using (_×_; <_,_>; uncurry) open import Relation.Nullary.Decidable.Core using (map′; _×-dec_; yes; no) open import Relation.Binary.Definitions using (DecidableEquality) @@ -66,10 +66,10 @@ function cs ≟ function cs′ = map′ (cong function) function-injective (cs Term.≟-Clauses cs′) data-type pars cs ≟ data-type pars′ cs′ = map′ (uncurry (cong₂ data-type)) data-type-injective - (pars ℕₚ.≟ pars′ ×-dec Listₚ.≡-dec Name._≟_ cs cs′) + (pars ℕ.≟ pars′ ×-dec List.≡-dec Name._≟_ cs cs′) record′ c fs ≟ record′ c′ fs′ = map′ (uncurry (cong₂ record′)) record′-injective - (c Name.≟ c′ ×-dec Listₚ.≡-dec (Arg.≡-dec Name._≟_) fs fs′) + (c Name.≟ c′ ×-dec List.≡-dec (Arg.≡-dec Name._≟_) fs fs′) constructor′ d ≟ constructor′ d′ = map′ (cong constructor′) constructor′-injective (d Name.≟ d′) axiom ≟ axiom = yes refl diff --git a/src/Reflection/AST/Meta.agda b/src/Reflection/AST/Meta.agda index babcd23670..a3cac80c9a 100644 --- a/src/Reflection/AST/Meta.agda +++ b/src/Reflection/AST/Meta.agda @@ -8,7 +8,7 @@ module Reflection.AST.Meta where -import Data.Nat.Properties as ℕₚ +import Data.Nat.Properties as ℕ open import Function.Base using (_on_) open import Relation.Nullary.Decidable.Core using (map′) open import Relation.Binary.Core using (Rel) @@ -30,7 +30,7 @@ _≈_ : Rel Meta _ _≈_ = _≡_ on toℕ _≈?_ : Decidable _≈_ -_≈?_ = On.decidable toℕ _≡_ ℕₚ._≟_ +_≈?_ = On.decidable toℕ _≡_ ℕ._≟_ _≟_ : DecidableEquality Meta m ≟ n = map′ (toℕ-injective _ _) (cong toℕ) (m ≈? n) diff --git a/src/Relation/Binary/Construct/Closure/Reflexive/Properties.agda b/src/Relation/Binary/Construct/Closure/Reflexive/Properties.agda index 29defea059..ff04ed09a1 100644 --- a/src/Relation/Binary/Construct/Closure/Reflexive/Properties.agda +++ b/src/Relation/Binary/Construct/Closure/Reflexive/Properties.agda @@ -8,7 +8,7 @@ module Relation.Binary.Construct.Closure.Reflexive.Properties where -open import Data.Product.Base as Prod +open import Data.Product.Base as Product open import Data.Sum.Base as Sum open import Function.Bundles using (_⇔_; mk⇔) open import Function.Base using (id) @@ -103,7 +103,7 @@ module _ {_~_ : Rel A ℓ} {P : Pred A p} where module _ {_~_ : Rel A ℓ} {P : Rel A p} where resp₂ : P Respects₂ _~_ → P Respects₂ (ReflClosure _~_) - resp₂ = Prod.map respˡ respʳ + resp₂ = Product.map respˡ respʳ ------------------------------------------------------------------------ -- Structures diff --git a/src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda b/src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda index a78d14938b..a68071d25e 100644 --- a/src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda +++ b/src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda @@ -14,10 +14,10 @@ open import Relation.Binary.Bundles using (Preorder) open import Relation.Binary.Structures using (IsPreorder) open import Relation.Binary.Definitions using (Transitive; Reflexive) open import Relation.Binary.Construct.Closure.ReflexiveTransitive -open import Relation.Binary.PropositionalEquality.Core as PropEq +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl; sym; cong; cong₂) -import Relation.Binary.PropositionalEquality.Properties as PropEq -import Relation.Binary.Reasoning.Preorder as PreorderReasoning +import Relation.Binary.PropositionalEquality.Properties as ≡ +import Relation.Binary.Reasoning.Preorder as ≲-Reasoning open import Relation.Binary.Reasoning.Syntax ------------------------------------------------------------------------ @@ -82,7 +82,7 @@ fold-◅◅ P _⊕_ ∅ left-unit assoc (x ◅ xs) ys = begin 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 + where open ≡.≡-Reasoning ------------------------------------------------------------------------ -- Relational properties @@ -97,7 +97,7 @@ module _ {i t} {I : Set i} (T : Rel I t) where isPreorder : IsPreorder _≡_ (Star T) isPreorder = record - { isEquivalence = PropEq.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = reflexive ; trans = trans } @@ -113,7 +113,7 @@ module _ {i t} {I : Set i} (T : Rel I t) where -- Preorder reasoning for Star module StarReasoning {i t} {I : Set i} (T : Rel I t) where - private module Base = PreorderReasoning (preorder T) + private module Base = ≲-Reasoning (preorder T) open Base public hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨; step-∼; step-≲) diff --git a/src/Relation/Binary/Construct/NaturalOrder/Left.agda b/src/Relation/Binary/Construct/NaturalOrder/Left.agda index c4846142ec..47465e8926 100644 --- a/src/Relation/Binary/Construct/NaturalOrder/Left.agda +++ b/src/Relation/Binary/Construct/NaturalOrder/Left.agda @@ -18,7 +18,7 @@ open import Relation.Binary.Structures open import Relation.Binary.Definitions using (Symmetric; Transitive; Reflexive; Antisymmetric; Total; _Respectsʳ_; _Respectsˡ_; _Respects₂_; Decidable) open import Relation.Nullary.Negation using (¬_) -import Relation.Binary.Reasoning.Setoid as EqReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning open import Relation.Binary.Lattice using (Infimum) module Relation.Binary.Construct.NaturalOrder.Left @@ -44,7 +44,7 @@ reflexive magma idem {x} {y} x≈y = begin x ≈⟨ sym (idem x) ⟩ x ∙ x ≈⟨ ∙-cong refl x≈y ⟩ x ∙ y ∎ - where open IsMagma magma; open EqReasoning setoid + where open IsMagma magma; open ≈-Reasoning setoid refl : Symmetric _≈_ → Idempotent _∙_ → Reflexive _≤_ refl sym idem {x} = sym (idem x) @@ -55,7 +55,7 @@ antisym isEq comm {x} {y} x≤y y≤x = begin x ∙ y ≈⟨ comm x y ⟩ y ∙ x ≈⟨ sym y≤x ⟩ y ∎ - where open IsEquivalence isEq; open EqReasoning (record { isEquivalence = isEq }) + where open IsEquivalence isEq; open ≈-Reasoning (record { isEquivalence = isEq }) total : Symmetric _≈_ → Transitive _≈_ → Selective _∙_ → Commutative _∙_ → Total _≤_ total sym trans sel comm x y with sel x y @@ -69,14 +69,14 @@ trans semi {x} {y} {z} x≤y y≤z = begin x ∙ (y ∙ z) ≈⟨ sym (assoc x y z) ⟩ (x ∙ y) ∙ z ≈⟨ ∙-cong (sym x≤y) S.refl ⟩ x ∙ z ∎ - where open module S = IsSemigroup semi; open EqReasoning S.setoid + where open module S = IsSemigroup semi; open ≈-Reasoning S.setoid respʳ : IsMagma _∙_ → _≤_ Respectsʳ _≈_ respʳ magma {x} {y} {z} y≈z x≤y = begin x ≈⟨ x≤y ⟩ x ∙ y ≈⟨ ∙-cong M.refl y≈z ⟩ x ∙ z ∎ - where open module M = IsMagma magma; open EqReasoning M.setoid + where open module M = IsMagma magma; open ≈-Reasoning M.setoid respˡ : IsMagma _∙_ → _≤_ Respectsˡ _≈_ respˡ magma {x} {y} {z} y≈z y≤x = begin @@ -84,7 +84,7 @@ respˡ magma {x} {y} {z} y≈z y≤x = begin y ≈⟨ y≤x ⟩ y ∙ x ≈⟨ ∙-cong y≈z M.refl ⟩ z ∙ x ∎ - where open module M = IsMagma magma; open EqReasoning M.setoid + where open module M = IsMagma magma; open ≈-Reasoning M.setoid resp₂ : IsMagma _∙_ → _≤_ Respects₂ _≈_ resp₂ magma = respʳ magma , respˡ magma @@ -95,7 +95,7 @@ dec _≟_ x y = x ≟ (x ∙ y) module _ (semi : IsSemilattice _∙_) where private open module S = IsSemilattice semi - open EqReasoning setoid + open ≈-Reasoning setoid x∙y≤x : ∀ x y → (x ∙ y) ≤ x x∙y≤x x y = begin diff --git a/src/Relation/Binary/Construct/NaturalOrder/Right.agda b/src/Relation/Binary/Construct/NaturalOrder/Right.agda index 3c2aa348d6..76529350e2 100644 --- a/src/Relation/Binary/Construct/NaturalOrder/Right.agda +++ b/src/Relation/Binary/Construct/NaturalOrder/Right.agda @@ -19,7 +19,7 @@ open import Relation.Binary.Definitions using (Symmetric; Transitive; Reflexive; Antisymmetric; Total; _Respectsʳ_; _Respectsˡ_; _Respects₂_; Decidable) open import Relation.Nullary.Negation using (¬_) -import Relation.Binary.Reasoning.Setoid as EqReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning module Relation.Binary.Construct.NaturalOrder.Right {a ℓ} {A : Set a} (_≈_ : Rel A ℓ) (_∙_ : Op₂ A) where @@ -44,7 +44,7 @@ reflexive magma idem {x} {y} x≈y = begin x ≈⟨ sym (idem x) ⟩ x ∙ x ≈⟨ ∙-cong x≈y refl ⟩ y ∙ x ∎ - where open IsMagma magma; open EqReasoning setoid + where open IsMagma magma; open ≈-Reasoning setoid refl : Symmetric _≈_ → Idempotent _∙_ → Reflexive _≤_ refl sym idem {x} = sym (idem x) @@ -55,7 +55,7 @@ antisym isEq comm {x} {y} x≤y y≤x = begin y ∙ x ≈⟨ comm y x ⟩ x ∙ y ≈⟨ y≤x ⟨ y ∎ - where open EqReasoning (record { isEquivalence = isEq }) + where open ≈-Reasoning (record { isEquivalence = isEq }) total : Symmetric _≈_ → Transitive _≈_ → Selective _∙_ → Commutative _∙_ → Total _≤_ total sym trans sel comm x y with sel x y @@ -69,14 +69,14 @@ trans semi {x} {y} {z} x≤y y≤z = begin (z ∙ y) ∙ x ≈⟨ assoc z y x ⟩ z ∙ (y ∙ x) ≈⟨ ∙-cong S.refl (sym x≤y) ⟩ z ∙ x ∎ - where open module S = IsSemigroup semi; open EqReasoning S.setoid + where open module S = IsSemigroup semi; open ≈-Reasoning S.setoid respʳ : IsMagma _∙_ → _≤_ Respectsʳ _≈_ respʳ magma {x} {y} {z} y≈z x≤y = begin x ≈⟨ x≤y ⟩ y ∙ x ≈⟨ ∙-cong y≈z M.refl ⟩ z ∙ x ∎ - where open module M = IsMagma magma; open EqReasoning M.setoid + where open module M = IsMagma magma; open ≈-Reasoning M.setoid respˡ : IsMagma _∙_ → _≤_ Respectsˡ _≈_ respˡ magma {x} {y} {z} y≈z y≤x = begin @@ -84,7 +84,7 @@ respˡ magma {x} {y} {z} y≈z y≤x = begin y ≈⟨ y≤x ⟩ x ∙ y ≈⟨ ∙-cong M.refl y≈z ⟩ x ∙ z ∎ - where open module M = IsMagma magma; open EqReasoning M.setoid + where open module M = IsMagma magma; open ≈-Reasoning M.setoid resp₂ : IsMagma _∙_ → _≤_ Respects₂ _≈_ resp₂ magma = respʳ magma , respˡ magma diff --git a/src/Relation/Binary/Construct/Subst/Equality.agda b/src/Relation/Binary/Construct/Subst/Equality.agda index 95172b57ee..8cde3b1a66 100644 --- a/src/Relation/Binary/Construct/Subst/Equality.agda +++ b/src/Relation/Binary/Construct/Subst/Equality.agda @@ -9,7 +9,7 @@ {-# OPTIONS --cubical-compatible --safe #-} -open import Data.Product.Base as Prod +open import Data.Product.Base using (_,_) open import Relation.Binary.Core using (Rel; _⇔_) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive) diff --git a/src/Relation/Binary/Definitions.agda b/src/Relation/Binary/Definitions.agda index 887cb7c4a6..dd0e016c12 100644 --- a/src/Relation/Binary/Definitions.agda +++ b/src/Relation/Binary/Definitions.agda @@ -12,14 +12,12 @@ module Relation.Binary.Definitions where open import Agda.Builtin.Equality using (_≡_) -open import Data.Maybe.Base using (Maybe) open import Data.Product.Base using (_×_; ∃-syntax) open import Data.Sum.Base using (_⊎_) open import Function.Base using (_on_; flip) open import Level open import Relation.Binary.Core -open import Relation.Nullary.Decidable.Core using (Dec) -open import Relation.Nullary.Negation.Core using (¬_) +open import Relation.Nullary as Nullary using (¬_; Dec) private variable @@ -206,35 +204,40 @@ P Respects₂ _∼_ = (P Respectsʳ _∼_) × (P Respectsˡ _∼_) Substitutive : Rel A ℓ₁ → (ℓ₂ : Level) → Set _ Substitutive {A = A} _∼_ p = (P : A → Set p) → P Respects _∼_ --- Decidability - it is possible to determine whether a given pair of --- elements are related. +-- Irrelevancy - all proofs that a given pair of elements are related +-- are indistinguishable. -Decidable : REL A B ℓ → Set _ -Decidable _∼_ = ∀ x y → Dec (x ∼ y) +Irrelevant : REL A B ℓ → Set _ +Irrelevant _∼_ = ∀ {x y} → Nullary.Irrelevant (x ∼ y) + +-- Recomputability - we can rebuild a relevant proof given an +-- irrelevant one. + +Recomputable : REL A B ℓ → Set _ +Recomputable _∼_ = ∀ {x y} → Nullary.Recomputable (x ∼ y) + +-- Stability + +Stable : REL A B ℓ → Set _ +Stable _∼_ = ∀ x y → Nullary.Stable (x ∼ y) -- Weak decidability - it is sometimes possible to determine if a given -- pair of elements are related. WeaklyDecidable : REL A B ℓ → Set _ -WeaklyDecidable _∼_ = ∀ x y → Maybe (x ∼ y) +WeaklyDecidable _∼_ = ∀ x y → Nullary.WeaklyDecidable (x ∼ y) + +-- Decidability - it is possible to determine whether a given pair of +-- elements are related. + +Decidable : REL A B ℓ → Set _ +Decidable _∼_ = ∀ x y → Dec (x ∼ y) -- Propositional equality is decidable for the type. DecidableEquality : (A : Set a) → Set _ DecidableEquality A = Decidable {A = A} _≡_ --- Irrelevancy - all proofs that a given pair of elements are related --- are indistinguishable. - -Irrelevant : REL A B ℓ → Set _ -Irrelevant _∼_ = ∀ {x y} (a b : x ∼ y) → a ≡ b - --- Recomputability - we can rebuild a relevant proof given an --- irrelevant one. - -Recomputable : REL A B ℓ → Set _ -Recomputable _∼_ = ∀ {x y} → .(x ∼ y) → x ∼ y - -- Universal - all pairs of elements are related Universal : REL A B ℓ → Set _ diff --git a/src/Relation/Binary/Lattice/Properties/HeytingAlgebra.agda b/src/Relation/Binary/Lattice/Properties/HeytingAlgebra.agda index e7090dc6c5..4727088718 100644 --- a/src/Relation/Binary/Lattice/Properties/HeytingAlgebra.agda +++ b/src/Relation/Binary/Lattice/Properties/HeytingAlgebra.agda @@ -19,13 +19,13 @@ open import Data.Product.Base using (_,_) open import Function.Base using (_$_; flip; _∘_) open import Level using (_⊔_) open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_) -import Relation.Binary.Reasoning.PartialOrder as POR +import Relation.Binary.Reasoning.PartialOrder as ≤-Reasoning open import Relation.Binary.Lattice.Properties.MeetSemilattice meetSemilattice open import Relation.Binary.Lattice.Properties.JoinSemilattice joinSemilattice import Relation.Binary.Lattice.Properties.BoundedMeetSemilattice boundedMeetSemilattice as BM open import Relation.Binary.Lattice.Properties.Lattice lattice open import Relation.Binary.Lattice.Properties.BoundedLattice boundedLattice -import Relation.Binary.Reasoning.Setoid as EqReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning ------------------------------------------------------------------------ -- Useful lemmas @@ -62,7 +62,7 @@ y≤x⇨y = transpose-⇨ (x∧y≤x _ _) x ⇨ u ≤⟨ ⇨ʳ-covariant u≤v ⟩ x ⇨ v ≤⟨ ⇨ˡ-contravariant y≤x ⟩ y ⇨ v ∎ - where open POR poset + where open ≤-Reasoning poset ⇨-cong : _⇨_ Preserves₂ _≈_ ⟶ _≈_ ⟶ _≈_ ⇨-cong x≈y u≈v = antisym (⇨-relax (reflexive $ Eq.sym x≈y) (reflexive u≈v)) @@ -113,7 +113,7 @@ y≤x⇨y = transpose-⇨ (x∧y≤x _ _) (((x ⇨ y) ∧ x) ∧ (x ⇨ z)) ∧ x ≈⟨ ∧-assoc _ _ _ ⟩ (((x ⇨ y) ∧ x) ∧ (x ⇨ z) ∧ x) ≤⟨ ∧-monotonic ⇨-eval ⇨-eval ⟩ y ∧ z ∎) - where open POR poset + where open ≤-Reasoning poset ⇨-distribˡ-∧ : _⇨_ DistributesOverˡ _∧_ ⇨-distribˡ-∧ x y z = antisym (⇨-distribˡ-∧-≤ x y z) (⇨-distribˡ-∧-≥ x y z) @@ -175,7 +175,7 @@ de-morgan₂-≤ x y = transpose-⇨ $ begin ¬ ¬ y ∧ ¬ y ≤⟨ ⇨-eval ⟩ ⊥ ∎ ⟩ ⊥ ∎ - where open POR poset + where open ≤-Reasoning poset de-morgan₂-≥ : ∀ x y → ¬ ¬ (¬ x ∨ ¬ y) ≤ ¬ (x ∧ y) de-morgan₂-≥ x y = transpose-⇨ $ ⇨-applyˡ $ transpose-⇨ $ begin @@ -184,7 +184,7 @@ de-morgan₂-≥ x y = transpose-⇨ $ ⇨-applyˡ $ transpose-⇨ $ begin (⇨-applyʳ (x∧y≤y _ _)) ⟩ ⊥ ∨ ⊥ ≈⟨ ∨-idempotent _ ⟩ ⊥ ∎ - where open POR poset + where open ≤-Reasoning poset de-morgan₂ : ∀ x y → ¬ (x ∧ y) ≈ ¬ ¬ (¬ x ∨ ¬ y) de-morgan₂ x y = antisym (de-morgan₂-≤ x y) (de-morgan₂-≥ x y) @@ -196,4 +196,4 @@ weak-lem {x} = begin ⊥ ∧ (x ⇨ ⊥) ⇨ ⊥ ≈⟨ ⇨-cong (∧-zeroˡ _) Eq.refl ⟩ ⊥ ⇨ ⊥ ≈⟨ ⇨-unit ⟩ ⊤ ∎ - where open EqReasoning setoid + where open ≈-Reasoning setoid diff --git a/src/Relation/Binary/Lattice/Properties/Lattice.agda b/src/Relation/Binary/Lattice/Properties/Lattice.agda index fd95effaaf..261371a3f0 100644 --- a/src/Relation/Binary/Lattice/Properties/Lattice.agda +++ b/src/Relation/Binary/Lattice/Properties/Lattice.agda @@ -21,8 +21,8 @@ open import Function.Base using (flip) open import Relation.Binary.Properties.Poset poset import Relation.Binary.Lattice.Properties.JoinSemilattice joinSemilattice as J import Relation.Binary.Lattice.Properties.MeetSemilattice meetSemilattice as M -import Relation.Binary.Reasoning.Setoid as EqReasoning -import Relation.Binary.Reasoning.PartialOrder as POR +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning +import Relation.Binary.Reasoning.PartialOrder as ≤-Reasoning ∨-absorbs-∧ : _∨_ Absorbs _∧_ ∨-absorbs-∧ x y = @@ -44,7 +44,7 @@ absorptive = ∨-absorbs-∧ , ∧-absorbs-∨ x ∧ y ≤⟨ x∧y≤x x y ⟩ x ≤⟨ x≤x∨y x y ⟩ x ∨ y ∎ - where open POR poset + where open ≤-Reasoning poset -- two quadrilateral arguments @@ -55,14 +55,14 @@ quadrilateral₁ {x} {y} x∨y≈x = begin y ∧ (x ∨ y) ≈⟨ M.∧-cong Eq.refl (J.∨-comm _ _) ⟩ y ∧ (y ∨ x) ≈⟨ ∧-absorbs-∨ _ _ ⟩ y ∎ - where open EqReasoning setoid + where open ≈-Reasoning setoid quadrilateral₂ : ∀ {x y} → x ∧ y ≈ y → x ∨ y ≈ x quadrilateral₂ {x} {y} x∧y≈y = begin x ∨ y ≈⟨ J.∨-cong Eq.refl (Eq.sym x∧y≈y) ⟩ x ∨ (x ∧ y) ≈⟨ ∨-absorbs-∧ _ _ ⟩ x ∎ - where open EqReasoning setoid + where open ≈-Reasoning setoid -- collapsing sublattice @@ -75,7 +75,7 @@ collapse₁ {x} {y} x≈y = begin x ∨ y ∎ where y≤x = reflexive (Eq.sym x≈y) - open EqReasoning setoid + open ≈-Reasoning setoid -- this can also be proved by quadrilateral argument, but it's much less symmetric. collapse₂ : ∀ {x y} → x ∨ y ≤ x ∧ y → x ≈ y @@ -88,7 +88,7 @@ collapse₂ {x} {y} ∨≤∧ = antisym x ∨ y ≤⟨ ∨≤∧ ⟩ x ∧ y ≤⟨ x∧y≤x _ _ ⟩ x ∎) - where open POR poset + where open ≤-Reasoning poset ------------------------------------------------------------------------ -- The dual construction is also a lattice. diff --git a/src/Relation/Binary/Properties/Preorder.agda b/src/Relation/Binary/Properties/Preorder.agda index b2be296828..93ff3ba71e 100644 --- a/src/Relation/Binary/Properties/Preorder.agda +++ b/src/Relation/Binary/Properties/Preorder.agda @@ -13,7 +13,7 @@ module Relation.Binary.Properties.Preorder {p₁ p₂ p₃} (P : Preorder p₁ p₂ p₃) where open import Function.Base using (flip) -open import Data.Product.Base as Prod using (_×_; _,_; swap) +open import Data.Product.Base as Product using (_×_; _,_; swap) import Relation.Binary.Construct.Flip.EqAndOrd as EqAndOrd open Preorder P @@ -37,7 +37,7 @@ InducedEquivalence = record ; isEquivalence = record { refl = (refl , refl) ; sym = swap - ; trans = Prod.zip trans (flip trans) + ; trans = Product.zip trans (flip trans) } } diff --git a/src/Relation/Binary/Reasoning/Base/Apartness.agda b/src/Relation/Binary/Reasoning/Base/Apartness.agda index 3d4d72f2c8..22eee2d505 100644 --- a/src/Relation/Binary/Reasoning/Base/Apartness.agda +++ b/src/Relation/Binary/Reasoning/Base/Apartness.agda @@ -13,7 +13,7 @@ open import Relation.Nullary.Decidable using (Dec; yes; no) open import Relation.Binary.Core using (Rel) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Transitive; Symmetric; Trans) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Reasoning.Syntax @@ -38,8 +38,8 @@ data _IsRelatedTo_ (x y : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where ≡-go : Trans _≡_ _IsRelatedTo_ _IsRelatedTo_ ≡-go x≡y nothing = nothing -≡-go x≡y (apartness y#z) = apartness (case x≡y of λ where P.refl → y#z) -≡-go x≡y (equals y≈z) = equals (case x≡y of λ where P.refl → y≈z) +≡-go x≡y (apartness y#z) = apartness (case x≡y of λ where ≡.refl → y#z) +≡-go x≡y (equals y≈z) = equals (case x≡y of λ where ≡.refl → y≈z) ≈-go : Trans _≈_ _IsRelatedTo_ _IsRelatedTo_ ≈-go x≈y nothing = nothing diff --git a/src/Relation/Binary/Reasoning/Base/Double.agda b/src/Relation/Binary/Reasoning/Base/Double.agda index 6fa901d573..b5fd9f66e4 100644 --- a/src/Relation/Binary/Reasoning/Base/Double.agda +++ b/src/Relation/Binary/Reasoning/Base/Double.agda @@ -16,7 +16,7 @@ open import Relation.Nullary.Decidable.Core using (Dec; yes; no) open import Relation.Binary.Core using (Rel; _⇒_) open import Relation.Binary.Definitions using (Reflexive; Trans) open import Relation.Binary.Structures using (IsPreorder) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Reasoning.Syntax @@ -40,8 +40,8 @@ start (equals x≈y) = reflexive x≈y start (nonstrict x≲y) = x≲y ≡-go : Trans _≡_ _IsRelatedTo_ _IsRelatedTo_ -≡-go x≡y (equals y≈z) = equals (case x≡y of λ where P.refl → y≈z) -≡-go x≡y (nonstrict y≤z) = nonstrict (case x≡y of λ where P.refl → y≤z) +≡-go x≡y (equals y≈z) = equals (case x≡y of λ where ≡.refl → y≈z) +≡-go x≡y (nonstrict y≤z) = nonstrict (case x≡y of λ where ≡.refl → y≤z) ≲-go : Trans _≲_ _IsRelatedTo_ _IsRelatedTo_ ≲-go x≲y (equals y≈z) = nonstrict (∼-respʳ-≈ y≈z x≲y) diff --git a/src/Relation/Binary/Reasoning/Base/Partial.agda b/src/Relation/Binary/Reasoning/Base/Partial.agda index bd1fcd8693..3d24edf501 100644 --- a/src/Relation/Binary/Reasoning/Base/Partial.agda +++ b/src/Relation/Binary/Reasoning/Base/Partial.agda @@ -11,7 +11,7 @@ open import Level using (_⊔_) open import Relation.Binary.Core open import Relation.Binary.Definitions open import Relation.Nullary.Decidable using (Dec; yes; no) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Reasoning.Syntax diff --git a/src/Relation/Binary/Reasoning/Base/Single.agda b/src/Relation/Binary/Reasoning/Base/Single.agda index ccbf63dda3..2f34ee8525 100644 --- a/src/Relation/Binary/Reasoning/Base/Single.agda +++ b/src/Relation/Binary/Reasoning/Base/Single.agda @@ -10,7 +10,7 @@ open import Level using (_⊔_) open import Function.Base using (case_of_) open import Relation.Binary.Core using (Rel; _⇒_) open import Relation.Binary.Definitions using (Reflexive; Transitive; Trans) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Reasoning.Syntax module Relation.Binary.Reasoning.Base.Single @@ -36,7 +36,7 @@ start (relTo x∼y) = x∼y ∼-go x∼y (relTo y∼z) = relTo (trans x∼y y∼z) ≡-go : Trans _≡_ _IsRelatedTo_ _IsRelatedTo_ -≡-go x≡y (relTo y∼z) = relTo (case x≡y of λ where P.refl → y∼z) +≡-go x≡y (relTo y∼z) = relTo (case x≡y of λ where ≡.refl → y∼z) stop : Reflexive _IsRelatedTo_ stop = relTo refl diff --git a/src/Relation/Binary/Reasoning/Base/Triple.agda b/src/Relation/Binary/Reasoning/Base/Triple.agda index 460115eff9..6afb937497 100644 --- a/src/Relation/Binary/Reasoning/Base/Triple.agda +++ b/src/Relation/Binary/Reasoning/Base/Triple.agda @@ -19,7 +19,7 @@ open import Relation.Binary.Core using (Rel; _⇒_) open import Relation.Binary.Structures using (IsPreorder) open import Relation.Binary.Definitions using (Transitive; _Respects₂_; Reflexive; Trans; Irreflexive; Asymmetric) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Reasoning.Syntax module Relation.Binary.Reasoning.Base.Triple {a ℓ₁ ℓ₂ ℓ₃} {A : Set a} @@ -48,9 +48,9 @@ start (nonstrict x≤y) = x≤y start (strict x_) open import Level using (Level; _⊔_; 0ℓ; suc; Lift) -open import Relation.Nullary.Decidable.Core using (Dec; True) -open import Relation.Nullary.Negation.Core using (¬_) +open import Relation.Nullary as Nullary using (¬_; Dec; True) open import Relation.Binary.PropositionalEquality.Core using (_≡_) private @@ -162,6 +161,28 @@ IUniversal P = ∀ {x} → x ∈ P syntax IUniversal P = ∀[ P ] +-- Irrelevance - any two proofs that an element satifies P are +-- indistinguishable. + +Irrelevant : Pred A ℓ → Set _ +Irrelevant P = ∀ {x} → Nullary.Irrelevant (P x) + +-- Recomputability - we can rebuild a relevant proof given an +-- irrelevant one. + +Recomputable : Pred A ℓ → Set _ +Recomputable P = ∀ {x} → Nullary.Recomputable (P x) + +-- Weak Decidability + +Stable : Pred A ℓ → Set _ +Stable P = ∀ x → Nullary.Stable (P x) + +-- Weak Decidability + +WeaklyDecidable : Pred A ℓ → Set _ +WeaklyDecidable P = ∀ x → Nullary.WeaklyDecidable (P x) + -- Decidability - it is possible to determine if an arbitrary element -- satisfies P. @@ -174,18 +195,6 @@ Decidable P = ∀ x → Dec (P x) ⌊_⌋ : {P : Pred A ℓ} → Decidable P → Pred A ℓ ⌊ P? ⌋ a = Lift _ (True (P? a)) --- Irrelevance - any two proofs that an element satifies P are --- indistinguishable. - -Irrelevant : Pred A ℓ → Set _ -Irrelevant P = ∀ {x} (a : P x) (b : P x) → a ≡ b - --- Recomputability - we can rebuild a relevant proof given an --- irrelevant one. - -Recomputable : Pred A ℓ → Set _ -Recomputable P = ∀ {x} → .(P x) → P x - ------------------------------------------------------------------------ -- Operations on sets diff --git a/src/System/Clock.agda b/src/System/Clock.agda index d8b8745b17..7ee4200207 100644 --- a/src/System/Clock.agda +++ b/src/System/Clock.agda @@ -14,7 +14,7 @@ open import Data.Fin.Base using (Fin; toℕ) open import Data.Nat.Base as ℕ using (ℕ; zero; suc; _+_; _∸_; _^_; _<ᵇ_) import Data.Nat.Show as ℕ using (show) open import Data.Nat.DivMod using (_/_) -import Data.Nat.Properties as ℕₚ +import Data.Nat.Properties as ℕ open import Data.String.Base using (String; _++_; padLeft) open import IO.Base @@ -104,5 +104,5 @@ show : Time → -- Time in seconds and nanoseconds show (mkTime s ns) prec = secs ++ "s" ++ padLeft '0' decimals nsecs where decimals = toℕ prec secs = ℕ.show s - prf = ℕₚ.m^n≢0 10 (9 ∸ decimals) + prf = ℕ.m^n≢0 10 (9 ∸ decimals) nsecs = ℕ.show ((ns / (10 ^ (9 ∸ decimals))) {{prf}}) diff --git a/src/Tactic/Cong.agda b/src/Tactic/Cong.agda index 2c1b4196b1..846fde2484 100644 --- a/src/Tactic/Cong.agda +++ b/src/Tactic/Cong.agda @@ -31,7 +31,7 @@ open import Data.Char.Base as Char using (toℕ) open import Data.Float.Base as Float using (_≡ᵇ_) open import Data.List.Base as List using ([]; _∷_) open import Data.Maybe.Base as Maybe using (Maybe; just; nothing) -open import Data.Nat.Base as Nat using (ℕ; zero; suc; _≡ᵇ_; _+_) +open import Data.Nat.Base as ℕ using (ℕ; zero; suc; _≡ᵇ_; _+_) open import Data.Unit.Base using (⊤) open import Data.Word.Base as Word using (toℕ) open import Data.Product.Base using (_×_; map₁; _,_) @@ -62,7 +62,7 @@ open import Reflection.TCM.Syntax private -- Descend past a variable. varDescend : ℕ → ℕ → ℕ - varDescend ϕ x = if ϕ Nat.≤ᵇ x then suc x else x + varDescend ϕ x = if ϕ ℕ.≤ᵇ x then suc x else x -- Descend a variable underneath pattern variables. patternDescend : ℕ → Pattern → Pattern × ℕ @@ -136,7 +136,7 @@ private antiUnifyClauses : ℕ → Clauses → Clauses → Maybe Clauses antiUnifyClause : ℕ → Clause → Clause → Maybe Clause - antiUnify ϕ (var x args) (var y args') with x Nat.≡ᵇ y | antiUnifyArgs ϕ args args' + antiUnify ϕ (var x args) (var y args') with x ℕ.≡ᵇ y | antiUnifyArgs ϕ args args' ... | _ | nothing = var ϕ [] ... | false | just uargs = var ϕ uargs ... | true | just uargs = var (varDescend ϕ x) uargs @@ -158,13 +158,13 @@ private Π[ s ∶ arg i (antiUnify ϕ a a') ] antiUnify (suc ϕ) b b' antiUnify ϕ (sort (set t)) (sort (set t')) = sort (set (antiUnify ϕ t t')) - antiUnify ϕ (sort (lit n)) (sort (lit n')) with n Nat.≡ᵇ n' + antiUnify ϕ (sort (lit n)) (sort (lit n')) with n ℕ.≡ᵇ n' ... | true = sort (lit n) ... | false = var ϕ [] - antiUnify ϕ (sort (propLit n)) (sort (propLit n')) with n Nat.≡ᵇ n' + antiUnify ϕ (sort (propLit n)) (sort (propLit n')) with n ℕ.≡ᵇ n' ... | true = sort (propLit n) ... | false = var ϕ [] - antiUnify ϕ (sort (inf n)) (sort (inf n')) with n Nat.≡ᵇ n' + antiUnify ϕ (sort (inf n)) (sort (inf n')) with n ℕ.≡ᵇ n' ... | true = sort (inf n) ... | false = var ϕ [] antiUnify ϕ (sort unknown) (sort unknown) = diff --git a/src/Tactic/MonoidSolver.agda b/src/Tactic/MonoidSolver.agda index c62ea9dccf..4e370bea25 100644 --- a/src/Tactic/MonoidSolver.agda +++ b/src/Tactic/MonoidSolver.agda @@ -88,7 +88,7 @@ import Reflection.AST.Name as Name open import Reflection.TCM open import Reflection.TCM.Syntax -import Relation.Binary.Reasoning.Setoid as SetoidReasoning +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning ------------------------------------------------------------------------ -- The Expr type with homomorphism proofs @@ -103,7 +103,7 @@ data Expr {a} (A : Set a) : Set a where module _ {m₁ m₂} (monoid : Monoid m₁ m₂) where open Monoid monoid - open SetoidReasoning setoid + open ≈-Reasoning setoid -- Convert the AST to an expression (i.e. evaluate it) without -- normalising. diff --git a/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Addition.agda b/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Addition.agda index bad78c4da2..e71ad07419 100644 --- a/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Addition.agda +++ b/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Addition.agda @@ -13,12 +13,12 @@ module Tactic.RingSolver.Core.Polynomial.Homomorphism.Addition (homo : Homomorphism r₁ r₂ r₃ r₄) where -open import Data.Nat as ℕ using (ℕ; suc; zero; compare; _≤′_; ≤′-step; ≤′-refl) -open import Data.Nat.Properties as ℕₚ using (≤′-trans) -open import Data.Product.Base using (_,_; _×_; proj₂) -open import Data.List.Base using (_∷_; []) +open import Data.Nat as ℕ using (ℕ; suc; zero; compare; _≤′_; ≤′-step; ≤′-refl) +open import Data.Nat.Properties as ℕ using (≤′-trans) +open import Data.Product.Base using (_,_; _×_; proj₂) +open import Data.List.Base using (_∷_; []) open import Data.List.Kleene -open import Data.Vec.Base using (Vec) +open import Data.Vec.Base using (Vec) open import Function.Base using (_⟨_⟩_; flip) open import Relation.Unary @@ -206,7 +206,7 @@ mutual ρ ^ suc j * ((ρ * ⅀?⟦ xs ⟧ (ρ , Ρ) + ⟦ poly x ⟧ Ρ) *⟨ ρ ⟩^ k) ≈⟨ pow-add _ _ k j ⟩ (ρ * ⅀?⟦ xs ⟧ (ρ , Ρ) + ⟦ poly x ⟧ Ρ) *⟨ ρ ⟩^ (k ℕ.+ suc j) - ≡⟨ ≡.cong (λ i → (ρ * ⅀?⟦ xs ⟧ (ρ , Ρ) + ⟦ poly x ⟧ Ρ) *⟨ ρ ⟩^ i) (ℕₚ.+-comm k (suc j)) ⟩ + ≡⟨ ≡.cong (λ i → (ρ * ⅀?⟦ xs ⟧ (ρ , Ρ) + ⟦ poly x ⟧ Ρ) *⟨ ρ ⟩^ i) (ℕ.+-comm k (suc j)) ⟩ (ρ * ⅀?⟦ xs ⟧ (ρ , Ρ) + ⟦ poly x ⟧ Ρ) *⟨ ρ ⟩^ (suc j ℕ.+ k) ∎) ⟩ diff --git a/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Exponentiation.agda b/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Exponentiation.agda index 4f6607d663..58086538ed 100644 --- a/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Exponentiation.agda +++ b/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Exponentiation.agda @@ -20,7 +20,7 @@ open import Data.Product.Base using (_,_; _×_; proj₁; proj₂) open import Data.List.Kleene open import Data.Vec.Base using (Vec) -import Data.Nat.Properties as ℕ-Prop +import Data.Nat.Properties as ℕ import Relation.Binary.PropositionalEquality.Core as ≡ open Homomorphism homo @@ -64,7 +64,7 @@ pow-eval-hom x (suc i) = (*-homo _ x) ⟨ trans ⟩ (≪* pow-eval-hom x i) rearrange zero = begin (ρ′ RawPow.^ (i ℕ.* 0)) * (⟦ x ⟧ Ρ RawPow.^ suc i) - ≡⟨ ≡.cong (λ k → (ρ′ RawPow.^ k) * (⟦ x ⟧ Ρ RawPow.^ suc i)) (ℕ-Prop.*-zeroʳ i) ⟩ + ≡⟨ ≡.cong (λ k → (ρ′ RawPow.^ k) * (⟦ x ⟧ Ρ RawPow.^ suc i)) (ℕ.*-zeroʳ i) ⟩ 1# * (⟦ x ⟧ Ρ RawPow.^ suc i) ≈⟨ *-identityˡ _ ⟩ ⟦ x ⟧ Ρ RawPow.^ suc i @@ -72,7 +72,7 @@ pow-eval-hom x (suc i) = (*-homo _ x) ⟨ trans ⟩ (≪* pow-eval-hom x i) rearrange j@(suc j′) = begin (ρ′ RawPow.^ (suc i ℕ.* j)) * (⟦ x ⟧ Ρ RawPow.^ suc i) - ≡⟨ ≡.cong (λ v → (ρ′ RawPow.^ v) * (⟦ x ⟧ Ρ RawPow.^ suc i)) (ℕ-Prop.*-comm (suc i) j) ⟩ + ≡⟨ ≡.cong (λ v → (ρ′ RawPow.^ v) * (⟦ x ⟧ Ρ RawPow.^ suc i)) (ℕ.*-comm (suc i) j) ⟩ (ρ′ RawPow.^ (j ℕ.* suc i)) * (⟦ x ⟧ Ρ RawPow.^ suc i) ≈⟨ ≪* sym (RawPow.^-assocʳ ρ′ j (suc i)) ⟩ ((ρ′ RawPow.^ suc j′) RawPow.^ suc i) * (⟦ x ⟧ Ρ RawPow.^ suc i) diff --git a/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Lemmas.agda b/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Lemmas.agda index f014d54c9f..c879bdef38 100644 --- a/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Lemmas.agda +++ b/src/Tactic/RingSolver/Core/Polynomial/Homomorphism/Lemmas.agda @@ -15,7 +15,7 @@ module Tactic.RingSolver.Core.Polynomial.Homomorphism.Lemmas open import Data.Bool using (Bool;true;false) open import Data.Nat.Base as ℕ using (ℕ; suc; zero; compare; _≤′_; ≤′-step; ≤′-refl) -open import Data.Nat.Properties as ℕₚ using (≤′-trans) +open import Data.Nat.Properties as ℕ using (≤′-trans) open import Data.Vec.Base as Vec using (Vec; _∷_) open import Data.Fin using (Fin; zero; suc) open import Data.List.Base using (_∷_; []) @@ -65,7 +65,7 @@ pow-hom : ∀ {n} i → (xs : Coeff n +) → ∀ ρ ρs → ⅀⟦ xs ⟧ (ρ , ρs) *⟨ ρ ⟩^ i ≈ ⅀⟦ xs ⍓+ i ⟧ (ρ , ρs) -pow-hom zero (x Δ j & xs) ρ ρs rewrite ℕₚ.+-identityʳ j = refl +pow-hom zero (x Δ j & xs) ρ ρs rewrite ℕ.+-identityʳ j = refl pow-hom (suc i) (x ≠0 Δ j & xs) ρ ρs = begin ρ ^ (suc i) * (((x , xs) ⟦∷⟧ (ρ , ρs)) *⟨ ρ ⟩^ j) diff --git a/src/Text/Pretty.agda b/src/Text/Pretty.agda index a4b3c1860d..a9202b1f82 100644 --- a/src/Text/Pretty.agda +++ b/src/Text/Pretty.agda @@ -26,8 +26,8 @@ open import Effect.Monad using (RawMonad) import Data.List.Effectful as List open RawMonad (List.monad {Level.zero}) -import Data.Nat.Properties as ℕₚ -open import Data.List.Extrema.Core ℕₚ.≤-totalOrder using (⊓ᴸ) +import Data.Nat.Properties as ℕ +open import Data.List.Extrema.Core ℕ.≤-totalOrder using (⊓ᴸ) ------------------------------------------------------------------------ -- Internal representation of documents and rendering function diff --git a/src/Text/Pretty/Core.agda b/src/Text/Pretty/Core.agda index 96415feb7f..39910f4dd5 100644 --- a/src/Text/Pretty/Core.agda +++ b/src/Text/Pretty/Core.agda @@ -23,15 +23,15 @@ import Data.Product.Relation.Unary.All as Allᴾ open import Data.Tree.Binary as Tree using (Tree; leaf; node; #nodes; mapₙ) open import Data.Tree.Binary.Relation.Unary.All as Allᵀ using (leaf; node) open import Data.Unit using (⊤; tt) -import Data.Tree.Binary.Relation.Unary.All.Properties as Allᵀₚ -import Data.Tree.Binary.Properties as Treeₚ +import Data.Tree.Binary.Relation.Unary.All.Properties as Allᵀ +import Data.Tree.Binary.Properties as Tree open import Data.Maybe.Base as Maybe using (Maybe; nothing; just; maybe′) open import Data.Maybe.Relation.Unary.All as Allᴹ using (nothing; just) open import Data.String.Base as String using (String; length; replicate; _++_; unlines) -open import Data.String.Unsafe as Stringₚ +open import Data.String.Unsafe as String open import Function.Base open import Relation.Nullary.Decidable using (Dec) open import Relation.Unary using (IUniversal; _⇒_; U) @@ -162,10 +162,10 @@ private size-indents : ∀ ma t → #nodes (indents ma t) ≡ #nodes t size-indents nothing t = refl - size-indents (just pad) t = Treeₚ.#nodes-mapₙ (pad ++_) t + size-indents (just pad) t = Tree.#nodes-mapₙ (pad ++_) t unfold-indents : ∀ ma t → indents ma t ≡ mapₙ (indent ma) t - unfold-indents nothing t = sym (Treeₚ.map-id t) + unfold-indents nothing t = sym (Tree.map-id t) unfold-indents (just pad) t = refl vContent : Content × String @@ -250,7 +250,7 @@ private All≤-node? (≤-Content (m≤m⊔n _ _) ∣xs∣) middle (subst (Allᵀ.All _ U) (sym $ unfold-indents pad tl) - $ Allᵀₚ.mapₙ⁺ (indent pad) (Allᵀ.mapₙ (indented _) ∣tl∣)) + $ Allᵀ.mapₙ⁺ (indent pad) (Allᵀ.mapₙ (indented _) ∣tl∣)) where middle : length (lastx ++ hd) ≤ vMaxWidth diff --git a/src/Text/Printf.agda b/src/Text/Printf.agda index dbb5a51d28..b7f1330c89 100644 --- a/src/Text/Printf.agda +++ b/src/Text/Printf.agda @@ -11,18 +11,17 @@ module Text.Printf where open import Data.String.Base using (String; fromChar; concat) open import Function.Base using (id) -import Data.Char.Base as Cₛ -import Data.Integer.Show as ℤₛ -import Data.Float as Fₛ -import Data.Nat.Show as ℕₛ using (show) +import Data.Integer.Show as ℤ +import Data.Float as Float +import Data.Nat.Show as ℕ open import Text.Format as Format hiding (Error) open import Text.Printf.Generic printfSpec : PrintfSpec formatSpec String -printfSpec .PrintfSpec.renderArg ℕArg = ℕₛ.show -printfSpec .PrintfSpec.renderArg ℤArg = ℤₛ.show -printfSpec .PrintfSpec.renderArg FloatArg = Fₛ.show +printfSpec .PrintfSpec.renderArg ℕArg = ℕ.show +printfSpec .PrintfSpec.renderArg ℤArg = ℤ.show +printfSpec .PrintfSpec.renderArg FloatArg = Float.show printfSpec .PrintfSpec.renderArg CharArg = fromChar printfSpec .PrintfSpec.renderArg StringArg = id printfSpec .PrintfSpec.renderStr = id diff --git a/src/Text/Regex/String.agda b/src/Text/Regex/String.agda index 911c130dc0..b526c77827 100644 --- a/src/Text/Regex/String.agda +++ b/src/Text/Regex/String.agda @@ -8,9 +8,9 @@ module Text.Regex.String where -import Data.Char.Properties as Charₚ +import Data.Char.Properties as Char ------------------------------------------------------------------------ -- Re-exporting definitions -open import Text.Regex Charₚ.≤-decPoset public +open import Text.Regex Char.≤-decPoset public diff --git a/src/Text/Tabular/List.agda b/src/Text/Tabular/List.agda index 1f1603424b..ca8cce5b70 100644 --- a/src/Text/Tabular/List.agda +++ b/src/Text/Tabular/List.agda @@ -10,7 +10,7 @@ module Text.Tabular.List where open import Data.String.Base using (String) open import Data.List.Base -import Data.Nat.Properties as ℕₚ +import Data.Nat.Properties as ℕ open import Data.Product.Base using (-,_; proj₂) open import Data.Vec.Base as Vec using (Vec) open import Data.Vec.Bounded.Base as Vec≤ using (Vec≤) @@ -24,7 +24,7 @@ display c a rows = Show.display c alignment rectangle where alignment : Vec Alignment _ alignment = Vec≤.padRight Left - $ Vec≤.≤-cast (ℕₚ.m⊓n≤m _ _) + $ Vec≤.≤-cast (ℕ.m⊓n≤m _ _) $ Vec≤.take _ (Vec≤.fromList a) rectangle : Vec (Vec String _) _