From 90deb34b59ad9f872ab898fef81d8f2d925ae2f4 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 13 Feb 2024 07:27:27 +0000 Subject: [PATCH 1/5] Qualified import of `PropositionalEquality` fixing #2280 --- .../Properties/BooleanAlgebra/Expression.agda | 16 +- src/Algebra/Properties/Monoid/Sum.agda | 6 +- src/Codata/Musical/Colist.agda | 22 +-- .../Colist/Relation/Unary/Any/Properties.agda | 12 +- src/Data/Container/Combinator/Properties.agda | 24 +-- src/Data/Container/Indexed.agda | 10 +- src/Data/Container/Indexed/Combinator.agda | 10 +- src/Data/Container/Morphism/Properties.agda | 6 +- .../Relation/Unary/Any/Properties.agda | 90 +++++------ src/Data/Fin/Permutation.agda | 10 +- src/Data/Fin/Properties.agda | 12 +- src/Data/Fin/Substitution/Example.agda | 2 +- src/Data/List/Effectful.agda | 32 ++-- .../List/Relation/Binary/Equality/Setoid.agda | 4 +- .../Binary/Sublist/Heterogeneous/Solver.agda | 8 +- .../List/Relation/Unary/First/Properties.agda | 12 +- .../Function/Dependent/Propositional.agda | 144 +++++++++--------- .../Dependent/Propositional/WithK.agda | 2 +- src/Function/Bijection.agda | 4 +- src/Function/Equivalence.agda | 4 +- src/Function/Injection.agda | 6 +- src/Function/Inverse.agda | 4 +- src/Function/LeftInverse.agda | 4 +- src/Function/Surjection.agda | 4 +- src/Relation/Binary/Structures.agda | 4 +- 25 files changed, 226 insertions(+), 226 deletions(-) diff --git a/src/Algebra/Lattice/Properties/BooleanAlgebra/Expression.agda b/src/Algebra/Lattice/Properties/BooleanAlgebra/Expression.agda index 64ac6d7864..5583f22a5b 100644 --- a/src/Algebra/Lattice/Properties/BooleanAlgebra/Expression.agda +++ b/src/Algebra/Lattice/Properties/BooleanAlgebra/Expression.agda @@ -26,7 +26,7 @@ open import Data.Vec.Properties using (lookup-map) open import Data.Vec.Relation.Binary.Pointwise.Extensional as PW using (Pointwise; ext) open import Function.Base using (_∘_; _$_; flip) -open import Relation.Binary.PropositionalEquality as P using (_≗_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) import Relation.Binary.Reflection as Reflection -- Expressions made up of variables and the operations of a boolean @@ -68,7 +68,7 @@ module Naturality (f : Applicative.Morphism A₁ A₂) where - open P.≡-Reasoning + open ≡.≡-Reasoning open Applicative.Morphism f open Semantics A₁ renaming (⟦_⟧ to ⟦_⟧₁) open Semantics A₂ renaming (⟦_⟧ to ⟦_⟧₂) @@ -77,21 +77,21 @@ module Naturality natural : ∀ {n} (e : Expr n) → op ∘ ⟦ e ⟧₁ ≗ ⟦ e ⟧₂ ∘ Vec.map op natural (var x) ρ = begin - op (Vec.lookup ρ x) ≡⟨ P.sym $ lookup-map x op ρ ⟩ + op (Vec.lookup ρ x) ≡⟨ ≡.sym $ lookup-map x op ρ ⟩ Vec.lookup (Vec.map op ρ) x ∎ natural (e₁ or e₂) ρ = begin op (_∨_ <$>₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩ - op (_∨_ <$>₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-<$> _ _) P.refl ⟩ - _∨_ <$>₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ (λ e₁ e₂ → _∨_ <$>₂ e₁ ⊛₂ e₂) (natural e₁ ρ) (natural e₂ ρ) ⟩ + op (_∨_ <$>₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ ≡.cong₂ _⊛₂_ (op-<$> _ _) ≡.refl ⟩ + _∨_ <$>₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ ≡.cong₂ (λ e₁ e₂ → _∨_ <$>₂ e₁ ⊛₂ e₂) (natural e₁ ρ) (natural e₂ ρ) ⟩ _∨_ <$>₂ ⟦ e₁ ⟧₂ (Vec.map op ρ) ⊛₂ ⟦ e₂ ⟧₂ (Vec.map op ρ) ∎ natural (e₁ and e₂) ρ = begin op (_∧_ <$>₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩ - op (_∧_ <$>₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-<$> _ _) P.refl ⟩ - _∧_ <$>₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ (λ e₁ e₂ → _∧_ <$>₂ e₁ ⊛₂ e₂) (natural e₁ ρ) (natural e₂ ρ) ⟩ + op (_∧_ <$>₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ ≡.cong₂ _⊛₂_ (op-<$> _ _) ≡.refl ⟩ + _∧_ <$>₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ ≡.cong₂ (λ e₁ e₂ → _∧_ <$>₂ e₁ ⊛₂ e₂) (natural e₁ ρ) (natural e₂ ρ) ⟩ _∧_ <$>₂ ⟦ e₁ ⟧₂ (Vec.map op ρ) ⊛₂ ⟦ e₂ ⟧₂ (Vec.map op ρ) ∎ natural (not e) ρ = begin op (¬_ <$>₁ ⟦ e ⟧₁ ρ) ≡⟨ op-<$> _ _ ⟩ - ¬_ <$>₂ op (⟦ e ⟧₁ ρ) ≡⟨ P.cong (¬_ <$>₂_) (natural e ρ) ⟩ + ¬_ <$>₂ op (⟦ e ⟧₁ ρ) ≡⟨ ≡.cong (¬_ <$>₂_) (natural e ρ) ⟩ ¬_ <$>₂ ⟦ e ⟧₂ (Vec.map op ρ) ∎ natural top ρ = begin op (pure₁ ⊤) ≡⟨ op-pure _ ⟩ diff --git a/src/Algebra/Properties/Monoid/Sum.agda b/src/Algebra/Properties/Monoid/Sum.agda index 6ba7d12fc8..fb383f3768 100644 --- a/src/Algebra/Properties/Monoid/Sum.agda +++ b/src/Algebra/Properties/Monoid/Sum.agda @@ -13,7 +13,7 @@ open import Data.Fin.Base using (zero; suc) open import Data.Unit using (tt) open import Function.Base using (_∘_) open import Relation.Binary.Core using (_Preserves_⟶_) -open import Relation.Binary.PropositionalEquality as P using (_≗_; _≡_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_; _≡_) module Algebra.Properties.Monoid.Sum {a ℓ} (M : Monoid a ℓ) where @@ -61,8 +61,8 @@ sum-cong-≋ {zero} xs≋ys = refl sum-cong-≋ {suc n} xs≋ys = +-cong (xs≋ys zero) (sum-cong-≋ (xs≋ys ∘ suc)) sum-cong-≗ : ∀ {n} → sum {n} Preserves _≗_ ⟶ _≡_ -sum-cong-≗ {zero} xs≗ys = P.refl -sum-cong-≗ {suc n} xs≗ys = P.cong₂ _+_ (xs≗ys zero) (sum-cong-≗ (xs≗ys ∘ suc)) +sum-cong-≗ {zero} xs≗ys = ≡.refl +sum-cong-≗ {suc n} xs≗ys = ≡.cong₂ _+_ (xs≗ys zero) (sum-cong-≗ (xs≗ys ∘ suc)) sum-replicate : ∀ n {x} → sum (replicate n x) ≈ n × x sum-replicate zero = refl diff --git a/src/Codata/Musical/Colist.agda b/src/Codata/Musical/Colist.agda index c65d33e8bf..78e8703be6 100644 --- a/src/Codata/Musical/Colist.agda +++ b/src/Codata/Musical/Colist.agda @@ -33,7 +33,7 @@ open import Relation.Binary.Definitions using (Transitive; Antisymmetric) import Relation.Binary.Construct.FromRel as Ind 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.PropositionalEquality as ≡ using (_≡_) open import Relation.Binary.Reasoning.Syntax open import Relation.Nullary.Reflects using (invert) open import Relation.Nullary @@ -107,23 +107,23 @@ Any-∈ {P = P} = mk↔ₛ′ from∘to where to : ∀ {xs} → Any P xs → ∃ λ x → x ∈ xs × P x - to (here p) = _ , here P.refl , p + to (here p) = _ , here ≡.refl , 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 + from (here ≡.refl) p = here p from (there x∈xs) p = there (from x∈xs p) to∘from : ∀ {x xs} (x∈xs : x ∈ xs) (p : P x) → to (from x∈xs p) ≡ (x , x∈xs , p) - to∘from (here P.refl) p = P.refl + to∘from (here ≡.refl) p = ≡.refl to∘from (there x∈xs) p = - P.cong (Product.map id (Product.map there id)) (to∘from x∈xs 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 - from∘to (here _) = P.refl - from∘to (there p) = P.cong there (from∘to p) + from∘to (here _) = ≡.refl + from∘to (there p) = ≡.cong there (from∘to p) -- Prefixes are subsets. @@ -176,8 +176,8 @@ module ⊑-Reasoning {a} {A : Set a} where ⊆-Preorder : ∀ {ℓ} → Set ℓ → Preorder _ _ _ ⊆-Preorder A = Ind.preorder (setoid A) _∈_ - (λ xs≈ys → ⊑⇒⊆ (⊑P.reflexive xs≈ys)) - where module ⊑P = Poset (⊑-Poset A) + (λ xs≈ys → ⊑⇒⊆ (⊑A.reflexive xs≈ys)) + where module ⊑A = Poset (⊑-Poset A) -- Example uses: -- @@ -220,7 +220,7 @@ infixr 5 _∷_ module Finite-injective where ∷-injective : ∀ {x : A} {xs p q} → (Finite (x ∷ xs) ∋ x ∷ p) ≡ x ∷ q → p ≡ q - ∷-injective P.refl = P.refl + ∷-injective ≡.refl = ≡.refl -- Infinite xs means that xs has infinite length. @@ -230,7 +230,7 @@ data Infinite {A : Set a} : Colist A → Set a where module Infinite-injective where ∷-injective : ∀ {x : A} {xs p q} → (Infinite (x ∷ xs) ∋ x ∷ p) ≡ x ∷ q → p ≡ q - ∷-injective P.refl = P.refl + ∷-injective ≡.refl = ≡.refl -- Colists which are not finite are infinite. diff --git a/src/Codata/Musical/Colist/Relation/Unary/Any/Properties.agda b/src/Codata/Musical/Colist/Relation/Unary/Any/Properties.agda index f7173c7e91..93c767f0b7 100644 --- a/src/Codata/Musical/Colist/Relation/Unary/Any/Properties.agda +++ b/src/Codata/Musical/Colist/Relation/Unary/Any/Properties.agda @@ -21,7 +21,7 @@ open import Function.Base using (_∋_; _∘_) open import Function.Bundles open import Level using (Level; _⊔_) open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality as P +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl; cong) open import Relation.Unary using (Pred) @@ -66,15 +66,15 @@ Any-cong {A = A} {P} {Q} {xs} {ys} P↔Q xs≈ys = to∘from : ∀ {xs ys} (xs≈ys : xs ≈ ys) (q : Any Q ys) → to xs≈ys (from xs≈ys q) ≡ q - to∘from (x ∷ xs≈) (there q) = P.cong there (to∘from (♭ xs≈) q) + to∘from (x ∷ xs≈) (there q) = ≡.cong there (to∘from (♭ xs≈) q) to∘from (x ∷ xs≈) (here qx) = - P.cong here (Inverse.strictlyInverseˡ P↔Q qx) + ≡.cong here (Inverse.strictlyInverseˡ P↔Q qx) from∘to : ∀ {xs ys} (xs≈ys : xs ≈ ys) (p : Any P xs) → from xs≈ys (to xs≈ys p) ≡ p - from∘to (x ∷ xs≈) (there p) = P.cong there (from∘to (♭ xs≈) p) + from∘to (x ∷ xs≈) (there p) = ≡.cong there (from∘to (♭ xs≈) p) from∘to (x ∷ xs≈) (here px) = - P.cong here (Inverse.strictlyInverseʳ P↔Q px) + ≡.cong here (Inverse.strictlyInverseʳ P↔Q px) ------------------------------------------------------------------------ -- map @@ -164,7 +164,7 @@ lookup-index (there p) = lookup-index p index-Any-resp : ∀ {f : ∀ {x} → P x → Q x} {xs ys} (xs≈ys : xs ≈ ys) (p : Any P xs) → index (Any-resp f xs≈ys p) ≡ index p -index-Any-resp (x ∷ xs≈) (here px) = P.refl +index-Any-resp (x ∷ xs≈) (here px) = ≡.refl index-Any-resp (x ∷ xs≈) (there p) = cong suc (index-Any-resp (♭ xs≈) p) diff --git a/src/Data/Container/Combinator/Properties.agda b/src/Data/Container/Combinator/Properties.agda index 6f7c83e993..5854138ec1 100644 --- a/src/Data/Container/Combinator/Properties.agda +++ b/src/Data/Container/Combinator/Properties.agda @@ -18,7 +18,7 @@ open import Data.Sum.Base as S using (inj₁; inj₂; [_,_]′; [_,_]) open import Function.Base as F using (_∘′_) open import Function.Bundles open import Level using (_⊔_; lower) -open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_) -- I have proved some of the correctness statements under the -- assumption of functional extensionality. I could have reformulated @@ -27,35 +27,35 @@ open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_) module Identity where correct : ∀ {s p x} {X : Set x} → ⟦ id {s} {p} ⟧ X ↔ F.id X - correct {X = X} = mk↔ₛ′ from-id to-id (λ _ → P.refl) (λ _ → P.refl) + correct {X = X} = mk↔ₛ′ from-id to-id (λ _ → ≡.refl) (λ _ → ≡.refl) module Constant (ext : ∀ {ℓ ℓ′} → Extensionality ℓ ℓ′) where correct : ∀ {x p y} (X : Set x) {Y : Set y} → ⟦ const {x} {p ⊔ y} X ⟧ Y ↔ F.const X Y - correct {x} {y} X {Y} = mk↔ₛ′ (from-const X) (to-const X) (λ _ → P.refl) from∘to + correct {x} {y} X {Y} = mk↔ₛ′ (from-const X) (to-const X) (λ _ → ≡.refl) from∘to where from∘to : (x : ⟦ const X ⟧ Y) → to-const X (proj₁ x) ≡ x - from∘to xs = P.cong (proj₁ xs ,_) (ext (λ x → ⊥-elim (lower x))) + from∘to xs = ≡.cong (proj₁ xs ,_) (ext (λ x → ⊥-elim (lower x))) module Composition {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) where correct : ∀ {x} {X : Set x} → ⟦ C₁ ∘ C₂ ⟧ X ↔ (⟦ C₁ ⟧ F.∘ ⟦ C₂ ⟧) X - correct {X = X} = mk↔ₛ′ (from-∘ C₁ C₂) (to-∘ C₁ C₂) (λ _ → P.refl) (λ _ → P.refl) + correct {X = X} = mk↔ₛ′ (from-∘ C₁ C₂) (to-∘ C₁ C₂) (λ _ → ≡.refl) (λ _ → ≡.refl) module Product (ext : ∀ {ℓ ℓ′} → Extensionality ℓ ℓ′) {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) where correct : ∀ {x} {X : Set x} → ⟦ C₁ × C₂ ⟧ X ↔ (⟦ C₁ ⟧ X Prod.× ⟦ C₂ ⟧ X) - correct {X = X} = mk↔ₛ′ (from-× C₁ C₂) (to-× C₁ C₂) (λ _ → P.refl) from∘to + correct {X = X} = mk↔ₛ′ (from-× C₁ C₂) (to-× C₁ C₂) (λ _ → ≡.refl) from∘to where from∘to : (to-× C₁ C₂) F.∘ (from-× C₁ C₂) ≗ F.id from∘to (s , f) = - P.cong (s ,_) (ext [ (λ _ → P.refl) , (λ _ → P.refl) ]) + ≡.cong (s ,_) (ext [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ]) module IndexedProduct {i s p} {I : Set i} (Cᵢ : I → Container s p) where correct : ∀ {x} {X : Set x} → ⟦ Π I Cᵢ ⟧ X ↔ (∀ i → ⟦ Cᵢ i ⟧ X) - correct {X = X} = mk↔ₛ′ (from-Π I Cᵢ) (to-Π I Cᵢ) (λ _ → P.refl) (λ _ → P.refl) + correct {X = X} = mk↔ₛ′ (from-Π I Cᵢ) (to-Π I Cᵢ) (λ _ → ≡.refl) (λ _ → ≡.refl) module Sum {s₁ s₂ p} (C₁ : Container s₁ p) (C₂ : Container s₂ p) where @@ -63,16 +63,16 @@ module Sum {s₁ s₂ p} (C₁ : Container s₁ p) (C₂ : Container s₂ p) whe correct {X = X} = mk↔ₛ′ (from-⊎ C₁ C₂) (to-⊎ C₁ C₂) to∘from from∘to where from∘to : (to-⊎ C₁ C₂) F.∘ (from-⊎ C₁ C₂) ≗ F.id - from∘to (inj₁ s₁ , f) = P.refl - from∘to (inj₂ s₂ , f) = P.refl + from∘to (inj₁ s₁ , f) = ≡.refl + from∘to (inj₂ s₂ , f) = ≡.refl to∘from : (from-⊎ C₁ C₂) F.∘ (to-⊎ C₁ C₂) ≗ F.id - to∘from = [ (λ _ → P.refl) , (λ _ → P.refl) ] + to∘from = [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] module IndexedSum {i s p} {I : Set i} (C : I → Container s p) where correct : ∀ {x} {X : Set x} → ⟦ Σ I C ⟧ X ↔ (∃ λ i → ⟦ C i ⟧ X) - correct {X = X} = mk↔ₛ′ (from-Σ I C) (to-Σ I C) (λ _ → P.refl) (λ _ → P.refl) + correct {X = X} = mk↔ₛ′ (from-Σ I C) (to-Σ I C) (λ _ → ≡.refl) (λ _ → ≡.refl) module ConstantExponentiation {i s p} {I : Set i} (C : Container s p) where diff --git a/src/Data/Container/Indexed.agda b/src/Data/Container/Indexed.agda index 65b36f0527..8507a2af39 100644 --- a/src/Data/Container/Indexed.agda +++ b/src/Data/Container/Indexed.agda @@ -18,7 +18,7 @@ open import Function.Base renaming (id to ⟨id⟩; _∘_ to _⟨∘⟩_) open import Function using (_↔_; Inverse) open import Relation.Unary using (Pred; _⊆_) open import Relation.Binary.Core using (Rel; REL) -open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_; refl) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_; refl) ------------------------------------------------------------------------ @@ -98,7 +98,7 @@ module _ {i₁ i₂ o₁ o₂} Container I₁ O₁ c₁ r → (I₁ → I₂) → (O₁ → O₂) → Container I₂ O₂ c₂ r → Set _ C₁ ⇒C[ f / g ] C₂ = ContainerMorphism C₁ C₂ f g _≡_ (λ R₂ R₁ → R₂ ≡ R₁) - (λ r₂≡r₁ r₂ → P.subst ⟨id⟩ r₂≡r₁ r₂) + (λ r₂≡r₁ r₂ → ≡.subst ⟨id⟩ r₂≡r₁ r₂) -- Degenerate cases where no reindexing is performed. @@ -123,7 +123,7 @@ module _ {i o c r} {I : Set i} {O : Set o} where ⟪_⟫ : ∀ {i o c r ℓ} {I : Set i} {O : Set o} {C₁ C₂ : Container I O c r} → C₁ ⇒ C₂ → (X : Pred I ℓ) → ⟦ C₁ ⟧ X ⊆ ⟦ C₂ ⟧ X ⟪ m ⟫ X (c , k) = command m c , λ r₂ → - P.subst X (coherent m) (k (response m r₂)) + ≡.subst X (coherent m) (k (response m r₂)) module PlainMorphism {i o c r} {I : Set i} {O : Set o} where @@ -145,7 +145,7 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where f ∘ g = record { command = command f ⟨∘⟩ command g ; response = response g ⟨∘⟩ response f - ; coherent = coherent g ⟨ P.trans ⟩ coherent f + ; coherent = coherent g ⟨ ≡.trans ⟩ coherent f } -- Identity commutes with ⟪_⟫. @@ -187,7 +187,7 @@ module CartesianMorphism morphism : C₁ ⇒ C₂ morphism = record { command = command m - ; response = P.subst ⟨id⟩ (response m) + ; response = ≡.subst ⟨id⟩ (response m) ; coherent = coherent m } diff --git a/src/Data/Container/Indexed/Combinator.agda b/src/Data/Container/Indexed/Combinator.agda index b35a1ec887..cb17210281 100644 --- a/src/Data/Container/Indexed/Combinator.agda +++ b/src/Data/Container/Indexed/Combinator.agda @@ -21,7 +21,7 @@ open import Function.Indexed.Bundles using (_↔ᵢ_) open import Level open import Relation.Unary using (Pred; _⊆_; _∪_; _∩_; ⋃; ⋂) renaming (_⟨×⟩_ to _⟪×⟫_; _⟨⊙⟩_ to _⟪⊙⟫_; _⟨⊎⟩_ to _⟪⊎⟫_) -open import Relation.Binary.PropositionalEquality as P +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_; refl) private @@ -167,7 +167,7 @@ module Constant (ext : ∀ {ℓ} → Extensionality ℓ ℓ) where from = < F.id , F.const ⊥-elim > to∘from : _ - to∘from xs = P.cong (proj₁ xs ,_) (ext ⊥-elim) + to∘from xs = ≡.cong (proj₁ xs ,_) (ext ⊥-elim) module Duality where @@ -202,7 +202,7 @@ module Product (ext : ∀ {ℓ} → Extensionality ℓ ℓ) where from∘to : from ⟨∘⟩ to ≗ F.id from∘to (c , _) = - P.cong (c ,_) (ext [ (λ _ → refl) , (λ _ → refl) ]) + ≡.cong (c ,_) (ext [ (λ _ → refl) , (λ _ → refl) ]) module IndexedProduct where @@ -231,8 +231,8 @@ module Sum (ext : ∀ {ℓ₁ ℓ₂} → Extensionality ℓ₁ ℓ₂) where from (inj₂ (c , f)) = inj₂ c , λ{ (All.inj₂ r) → f r} from∘to : from ⟨∘⟩ to ≗ F.id - from∘to (inj₁ _ , _) = P.cong (inj₁ _ ,_) (ext λ{ (All.inj₁ r) → refl}) - from∘to (inj₂ _ , _) = P.cong (inj₂ _ ,_) (ext λ{ (All.inj₂ r) → refl}) + from∘to (inj₁ _ , _) = ≡.cong (inj₁ _ ,_) (ext λ{ (All.inj₁ r) → refl}) + from∘to (inj₂ _ , _) = ≡.cong (inj₂ _ ,_) (ext λ{ (All.inj₂ r) → refl}) to∘from : to ⟨∘⟩ from ≗ F.id to∘from = [ (λ _ → refl) , (λ _ → refl) ] diff --git a/src/Data/Container/Morphism/Properties.agda b/src/Data/Container/Morphism/Properties.agda index c0d84007b4..39471b5e48 100644 --- a/src/Data/Container/Morphism/Properties.agda +++ b/src/Data/Container/Morphism/Properties.agda @@ -12,7 +12,7 @@ open import Level using (_⊔_; suc) open import Function.Base as F using (_$_) open import Data.Product.Base using (∃; proj₁; proj₂; _,_) open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_) open import Data.Container.Core open import Data.Container.Morphism @@ -23,7 +23,7 @@ open import Data.Container.Relation.Binary.Equality.Setoid module _ {s p} (C : Container s p) where id-correct : ∀ {x} {X : Set x} → ⟪ id C ⟫ {X = X} ≗ F.id - id-correct x = P.refl + id-correct x = ≡.refl -- Composition. @@ -33,7 +33,7 @@ module _ {s₁ s₂ s₃ p₁ p₂ p₃} ∘-correct : (f : C₂ ⇒ C₃) (g : C₁ ⇒ C₂) → ∀ {x} {X : Set x} → ⟪ f ∘ g ⟫ {X = X} ≗ (⟪ f ⟫ F.∘ ⟪ g ⟫) - ∘-correct f g xs = P.refl + ∘-correct f g xs = ≡.refl module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂} where diff --git a/src/Data/Container/Relation/Unary/Any/Properties.agda b/src/Data/Container/Relation/Unary/Any/Properties.agda index 0a4340fa7b..42ddeb6696 100644 --- a/src/Data/Container/Relation/Unary/Any/Properties.agda +++ b/src/Data/Container/Relation/Unary/Any/Properties.agda @@ -23,7 +23,7 @@ open import Function.Related.Propositional as Related using (Related; SK-sym) open import Function.Related.TypeIsomorphisms open import Relation.Unary using (Pred ; _∪_ ; _∩_) open import Relation.Binary.Core using (REL) -open import Relation.Binary.PropositionalEquality as P +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_; refl) private @@ -42,21 +42,21 @@ module _ {s p} (C : Container s p) {x} {X : Set x} {ℓ} {P : Pred X ℓ} where -- ◇ can be unwrapped to reveal the Σ type ↔Σ : ∀ {xs : ⟦ C ⟧ X} → ◇ C P xs ↔ ∃ λ p → P (proj₂ xs p) - ↔Σ {xs} = mk↔ₛ′ ◇.proof any (λ _ → P.refl) (λ _ → P.refl) + ↔Σ {xs} = mk↔ₛ′ ◇.proof any (λ _ → ≡.refl) (λ _ → ≡.refl) -- ◇ can be expressed using _∈_. ↔∈ : ∀ {xs : ⟦ C ⟧ X} → ◇ C P xs ↔ (∃ λ x → x ∈ xs × P x) - ↔∈ {xs} = mk↔ₛ′ to from to∘from (λ _ → P.refl) where + ↔∈ {xs} = mk↔ₛ′ to from to∘from (λ _ → ≡.refl) where to : ◇ C P xs → ∃ λ x → x ∈ xs × P x - to (any (p , Px)) = (proj₂ xs p , (any (p , P.refl)) , Px) + to (any (p , Px)) = (proj₂ xs p , (any (p , ≡.refl)) , Px) from : (∃ λ x → x ∈ xs × P x) → ◇ C P xs from (.(proj₂ xs p) , (any (p , refl)) , Px) = any (p , Px) to∘from : to ∘ from ≗ id - to∘from (.(proj₂ xs p) , any (p , refl) , Px) = P.refl + to∘from (.(proj₂ xs p) , any (p , refl) , Px) = ≡.refl module _ {s p} {C : Container s p} {x} {X : Set x} {ℓ₁ ℓ₂} {P₁ : Pred X ℓ₁} {P₂ : Pred X ℓ₂} where @@ -105,7 +105,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s flatten : ∀ (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) → ◇ C₁ (◇ C₂ P) xss ↔ ◇ (C₁ C.∘ C₂) P (Inverse.from (Composition.correct C₁ C₂) xss) - flatten xss = mk↔ₛ′ t f (λ _ → P.refl) (λ _ → P.refl) where + flatten xss = mk↔ₛ′ t f (λ _ → ≡.refl) (λ _ → ≡.refl) where ◇₁ = ◇ C₁; ◇₂ = ◇ C₂; ◇₁₂ = ◇ (C₁ C.∘ C₂) open Inverse @@ -132,11 +132,11 @@ module _ {s p} {C : Container s p} {x} {X : Set x} from = [ Any.map₂ inj₁ , Any.map₂ inj₂ ] from∘to : from ∘ to ≗ id - from∘to (any (pos , inj₁ p)) = P.refl - from∘to (any (pos , inj₂ q)) = P.refl + from∘to (any (pos , inj₁ p)) = ≡.refl + from∘to (any (pos , inj₂ q)) = ≡.refl to∘from : to ∘ from ≗ id - to∘from = [ (λ _ → P.refl) , (λ _ → P.refl) ] + to∘from = [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] -- Products "commute" with ◇. @@ -145,7 +145,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s ×◇↔◇◇× : ∀ {xs : ⟦ C₁ ⟧ X} {ys : ⟦ C₂ ⟧ Y} → ◇ C₁ (λ x → ◇ C₂ (λ y → P x × Q y) ys) xs ↔ (◇ C₁ P xs × ◇ C₂ Q ys) - ×◇↔◇◇× {xs} {ys} = mk↔ₛ′ to from (λ _ → P.refl) (λ _ → P.refl) + ×◇↔◇◇× {xs} {ys} = mk↔ₛ′ to from (λ _ → ≡.refl) (λ _ → ≡.refl) where ◇₁ = ◇ C₁; ◇₂ = ◇ C₂ @@ -209,7 +209,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s remove-linear {xs} m = mk↔ₛ′ t f t∘f f∘t where open _≃_ - open P.≡-Reasoning + open ≡.≡-Reasoning position⊸m : ∀ {s} → Position C₂ (shape⊸ m s) ≃ Position C₁ s position⊸m = ↔⇒≃ (position⊸ m) @@ -222,66 +222,66 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s f : ◇₁ P xs → ◇₂ P (⟪ m ⟫⊸ xs) f (any (x , p)) = any $ from position⊸m x - , P.subst (P ∘′ proj₂ xs) (P.sym (right-inverse-of position⊸m _)) p + , ≡.subst (P ∘′ proj₂ xs) (≡.sym (right-inverse-of position⊸m _)) p f∘t : f ∘ t ≗ id - f∘t (any (p₂ , p)) = P.cong any $ Σ-≡,≡→≡ + f∘t (any (p₂ , p)) = ≡.cong any $ Σ-≡,≡→≡ ( left-inverse-of position⊸m p₂ - , (P.subst (P ∘ proj₂ xs ∘ to position⊸m) + , (≡.subst (P ∘ proj₂ xs ∘ to position⊸m) (left-inverse-of position⊸m p₂) - (P.subst (P ∘ proj₂ xs) - (P.sym (right-inverse-of position⊸m + (≡.subst (P ∘ proj₂ xs) + (≡.sym (right-inverse-of position⊸m (to position⊸m p₂))) - p) ≡⟨ P.subst-∘ (left-inverse-of position⊸m _) ⟩ + p) ≡⟨ ≡.subst-∘ (left-inverse-of position⊸m _) ⟩ - P.subst (P ∘ proj₂ xs) - (P.cong (to position⊸m) + ≡.subst (P ∘ proj₂ xs) + (≡.cong (to position⊸m) (left-inverse-of position⊸m p₂)) - (P.subst (P ∘ proj₂ xs) - (P.sym (right-inverse-of position⊸m + (≡.subst (P ∘ proj₂ xs) + (≡.sym (right-inverse-of position⊸m (to position⊸m p₂))) - p) ≡⟨ P.cong (λ eq → P.subst (P ∘ proj₂ xs) eq - (P.subst (P ∘ proj₂ xs) - (P.sym (right-inverse-of position⊸m _)) _)) + p) ≡⟨ ≡.cong (λ eq → ≡.subst (P ∘ proj₂ xs) eq + (≡.subst (P ∘ proj₂ xs) + (≡.sym (right-inverse-of position⊸m _)) _)) (_≃_.left-right position⊸m _) ⟩ - P.subst (P ∘ proj₂ xs) + ≡.subst (P ∘ proj₂ xs) (right-inverse-of position⊸m (to position⊸m p₂)) - (P.subst (P ∘ proj₂ xs) - (P.sym (right-inverse-of position⊸m + (≡.subst (P ∘ proj₂ xs) + (≡.sym (right-inverse-of position⊸m (to position⊸m p₂))) - p) ≡⟨ P.subst-subst (P.sym (right-inverse-of position⊸m _)) ⟩ + p) ≡⟨ ≡.subst-subst (≡.sym (right-inverse-of position⊸m _)) ⟩ - P.subst (P ∘ proj₂ xs) - (P.trans - (P.sym (right-inverse-of position⊸m + ≡.subst (P ∘ proj₂ xs) + (≡.trans + (≡.sym (right-inverse-of position⊸m (to position⊸m p₂))) (right-inverse-of position⊸m (to position⊸m p₂))) - p ≡⟨ P.cong (λ eq → P.subst (P ∘ proj₂ xs) eq p) - (P.trans-symˡ (right-inverse-of position⊸m _)) ⟩ + p ≡⟨ ≡.cong (λ eq → ≡.subst (P ∘ proj₂ xs) eq p) + (≡.trans-symˡ (right-inverse-of position⊸m _)) ⟩ - P.subst (P ∘ proj₂ xs) P.refl p ≡⟨⟩ + ≡.subst (P ∘ proj₂ xs) ≡.refl p ≡⟨⟩ p ∎) ) t∘f : t ∘ f ≗ id - t∘f (any (p₁ , p)) = P.cong any $ Σ-≡,≡→≡ + t∘f (any (p₁ , p)) = ≡.cong any $ Σ-≡,≡→≡ ( right-inverse-of position⊸m p₁ - , (P.subst (P ∘ proj₂ xs) + , (≡.subst (P ∘ proj₂ xs) (right-inverse-of position⊸m p₁) - (P.subst (P ∘ proj₂ xs) - (P.sym (right-inverse-of position⊸m p₁)) - p) ≡⟨ P.subst-subst (P.sym (right-inverse-of position⊸m _)) ⟩ + (≡.subst (P ∘ proj₂ xs) + (≡.sym (right-inverse-of position⊸m p₁)) + p) ≡⟨ ≡.subst-subst (≡.sym (right-inverse-of position⊸m _)) ⟩ - P.subst (P ∘ proj₂ xs) - (P.trans - (P.sym (right-inverse-of position⊸m p₁)) + ≡.subst (P ∘ proj₂ xs) + (≡.trans + (≡.sym (right-inverse-of position⊸m p₁)) (right-inverse-of position⊸m p₁)) - p ≡⟨ P.cong (λ eq → P.subst (P ∘ proj₂ xs) eq p) - (P.trans-symˡ (right-inverse-of position⊸m _)) ⟩ - P.subst (P ∘ proj₂ xs) P.refl p ≡⟨⟩ + p ≡⟨ ≡.cong (λ eq → ≡.subst (P ∘ proj₂ xs) eq p) + (≡.trans-symˡ (right-inverse-of position⊸m _)) ⟩ + ≡.subst (P ∘ proj₂ xs) ≡.refl p ≡⟨⟩ p ∎) ) diff --git a/src/Data/Fin/Permutation.agda b/src/Data/Fin/Permutation.agda index 59a5c6abdf..96a557a4c7 100644 --- a/src/Data/Fin/Permutation.agda +++ b/src/Data/Fin/Permutation.agda @@ -28,9 +28,9 @@ open import Relation.Binary.Core using (Rel) open import Relation.Nullary using (does; ¬_; yes; no) open import Relation.Nullary.Decidable using (dec-yes; dec-no) open import Relation.Nullary.Negation using (contradiction) -open import Relation.Binary.PropositionalEquality as P +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≢_; refl; trans; sym; →-to-⟶; cong; cong₂) -open P.≡-Reasoning +open ≡.≡-Reasoning private variable @@ -67,10 +67,10 @@ _⟨$⟩ˡ_ : Permutation m n → Fin n → Fin m _⟨$⟩ˡ_ = Inverse.from inverseˡ : ∀ (π : Permutation m n) {i} → π ⟨$⟩ˡ (π ⟨$⟩ʳ i) ≡ i -inverseˡ π = Inverse.inverseʳ π P.refl +inverseˡ π = Inverse.inverseʳ π ≡.refl inverseʳ : ∀ (π : Permutation m n) {i} → π ⟨$⟩ʳ (π ⟨$⟩ˡ i) ≡ i -inverseʳ π = Inverse.inverseˡ π P.refl +inverseʳ π = Inverse.inverseˡ π ≡.refl ------------------------------------------------------------------------ -- Equality @@ -251,7 +251,7 @@ module _ (π : Permutation (suc m) (suc n)) where ↔⇒≡ {suc m} {suc n} π = cong suc (↔⇒≡ (remove 0F π)) fromPermutation : Permutation m n → Permutation′ m -fromPermutation π = P.subst (Permutation _) (sym (↔⇒≡ π)) π +fromPermutation π = ≡.subst (Permutation _) (sym (↔⇒≡ π)) π refute : m ≢ n → ¬ Permutation m n refute m≢n π = contradiction (↔⇒≡ π) m≢n diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda index d105d9e8d7..ff9239359c 100644 --- a/src/Data/Fin/Properties.agda +++ b/src/Data/Fin/Properties.agda @@ -41,7 +41,7 @@ open import Relation.Binary.Bundles using (Preorder; Setoid; DecSetoid; Poset; TotalOrder; DecTotalOrder; StrictPartialOrder; StrictTotalOrder) open import Relation.Binary.Structures using (IsDecEquivalence; IsPreorder; IsPartialOrder; IsTotalOrder; IsDecTotalOrder; IsStrictPartialOrder; IsStrictTotalOrder) -open import Relation.Binary.PropositionalEquality as P +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≢_; refl; sym; trans; cong; cong₂; subst; _≗_; module ≡-Reasoning) open import Relation.Nullary.Decidable as Dec using (Dec; _because_; yes; no; _×-dec_; _⊎-dec_; map′) @@ -104,7 +104,7 @@ suc x ≟ suc y = map′ (cong suc) suc-injective (x ≟ y) ≡-isDecEquivalence : IsDecEquivalence {A = Fin n} _≡_ ≡-isDecEquivalence = record - { isEquivalence = P.isEquivalence + { isEquivalence = ≡.isEquivalence ; _≟_ = _≟_ } @@ -112,10 +112,10 @@ suc x ≟ suc y = map′ (cong suc) suc-injective (x ≟ y) -- Bundles ≡-preorder : ℕ → Preorder _ _ _ -≡-preorder n = P.preorder (Fin n) +≡-preorder n = ≡.preorder (Fin n) ≡-setoid : ℕ → Setoid _ _ -≡-setoid n = P.setoid (Fin n) +≡-setoid n = ≡.setoid (Fin n) ≡-decSetoid : ℕ → DecSetoid _ _ ≡-decSetoid n = record @@ -314,7 +314,7 @@ m >= pure) ≡ xs right-identity [] = refl - right-identity (x ∷ xs) = P.cong (x ∷_) (right-identity xs) + right-identity (x ∷ xs) = ≡.cong (x ∷_) (right-identity xs) left-zero : ∀ {ℓ} {A B : Set ℓ} (f : A → List B) → (∅ >>= f) ≡ ∅ left-zero f = refl @@ -147,8 +147,8 @@ module MonadProperties where (xs ∣ ys >>= f) ≡ ((xs >>= f) ∣ (ys >>= f)) right-distributive [] ys f = refl right-distributive (x ∷ xs) ys f = begin - f x ∣ (xs ∣ ys >>= f) ≡⟨ P.cong (f x ∣_) $ right-distributive xs ys f ⟩ - f x ∣ ((xs >>= f) ∣ (ys >>= f)) ≡⟨ P.sym $ ++-assoc (f x) _ _ ⟩ + f x ∣ (xs ∣ ys >>= f) ≡⟨ ≡.cong (f x ∣_) $ right-distributive xs ys f ⟩ + f x ∣ ((xs >>= f) ∣ (ys >>= f)) ≡⟨ ≡.sym $ ++-assoc (f x) _ _ ⟩ ((f x ∣ (xs >>= f)) ∣ (ys >>= f)) ∎ associative : ∀ {ℓ} {A B C : Set ℓ} @@ -156,13 +156,13 @@ module MonadProperties where (xs >>= λ x → f x >>= g) ≡ (xs >>= f >>= g) associative [] f g = refl associative (x ∷ xs) f g = begin - (f x >>= g) ∣ (xs >>= λ x → f x >>= g) ≡⟨ P.cong ((f x >>= g) ∣_) $ associative xs f g ⟩ - (f x >>= g) ∣ (xs >>= f >>= g) ≡⟨ P.sym $ right-distributive (f x) (xs >>= f) g ⟩ + (f x >>= g) ∣ (xs >>= λ x → f x >>= g) ≡⟨ ≡.cong ((f x >>= g) ∣_) $ associative xs f g ⟩ + (f x >>= g) ∣ (xs >>= f >>= g) ≡⟨ ≡.sym $ right-distributive (f x) (xs >>= f) g ⟩ (f x ∣ (xs >>= f) >>= g) ∎ cong : ∀ {ℓ} {A B : Set ℓ} {xs₁ xs₂} {f₁ f₂ : A → List B} → xs₁ ≡ xs₂ → f₁ ≗ f₂ → (xs₁ >>= f₁) ≡ (xs₂ >>= f₂) - cong {xs₁ = xs} refl f₁≗f₂ = P.cong concat (map-cong f₁≗f₂ xs) + cong {xs₁ = xs} refl f₁≗f₂ = ≡.cong concat (map-cong f₁≗f₂ xs) ------------------------------------------------------------------------ -- The applicative functor derived from the list monad. @@ -202,14 +202,14 @@ module Applicative where unfold-<$> : ∀ {ℓ} {A B : Set ℓ} → (f : A → B) (as : List A) → (f <$> as) ≡ (pure f ⊛ as) - unfold-<$> f as = P.sym (++-identityʳ (f <$> as)) + unfold-<$> f as = ≡.sym (++-identityʳ (f <$> as)) -- _⊛_ unfolds to binds. unfold-⊛ : ∀ {ℓ} {A B : Set ℓ} → (fs : List (A → B)) (as : List A) → (fs ⊛ as) ≡ (fs >>= pam as) unfold-⊛ fs as = begin fs ⊛ as - ≡⟨ concatMap-cong (λ f → P.cong (map f) (concatMap-pure as)) fs ⟨ + ≡⟨ concatMap-cong (λ f → ≡.cong (map f) (concatMap-pure as)) fs ⟨ concatMap (λ f → map f (concatMap pure as)) fs ≡⟨ concatMap-cong (λ f → map-concatMap f pure as) fs ⟩ concatMap (λ f → concatMap (λ x → pure (f x)) as) fs @@ -224,7 +224,7 @@ module Applicative where right-distributive fs₁ fs₂ xs = begin (fs₁ ∣ fs₂) ⊛ xs ≡⟨ unfold-⊛ (fs₁ ∣ fs₂) xs ⟩ (fs₁ ∣ fs₂ >>= pam xs) ≡⟨ MonadProperties.right-distributive fs₁ fs₂ (pam xs) ⟩ - (fs₁ >>= pam xs) ∣ (fs₂ >>= pam xs) ≡⟨ P.cong₂ _∣_ (unfold-⊛ fs₁ xs) (unfold-⊛ fs₂ xs) ⟨ + (fs₁ >>= pam xs) ∣ (fs₂ >>= pam xs) ≡⟨ ≡.cong₂ _∣_ (unfold-⊛ fs₁ xs) (unfold-⊛ fs₂ xs) ⟨ (fs₁ ⊛ xs ∣ fs₂ ⊛ xs) ∎ -- _⊛_ does not distribute over _∣_ from the left. @@ -262,9 +262,9 @@ module Applicative where pure _∘′_ ⊛ fs ⊛ gs ⊛ xs ≡⟨ unfold-⊛ (pure _∘′_ ⊛ fs ⊛ gs) xs ⟩ (pure _∘′_ ⊛ fs ⊛ gs >>= pam xs) - ≡⟨ P.cong (_>>= pam xs) (unfold-⊛ (pure _∘′_ ⊛ fs) gs) ⟩ + ≡⟨ ≡.cong (_>>= pam xs) (unfold-⊛ (pure _∘′_ ⊛ fs) gs) ⟩ (pure _∘′_ ⊛ fs >>= pam gs >>= pam xs) - ≡⟨ P.cong (λ h → h >>= pam gs >>= pam xs) (unfold-⊛ (pure _∘′_) fs) ⟩ + ≡⟨ ≡.cong (λ h → h >>= pam gs >>= pam xs) (unfold-⊛ (pure _∘′_) fs) ⟩ (pure _∘′_ >>= pam fs >>= pam gs >>= pam xs) ≡⟨ MP.cong (MP.cong (MP.left-identity _∘′_ (pam fs)) (λ f → refl {x = pam gs f})) @@ -278,13 +278,13 @@ module Applicative where (fs >>= λ f → gs >>= λ g → pam xs (f ∘′ g)) ≡⟨ (MP.cong (refl {x = fs}) λ f → MP.cong (refl {x = gs}) λ g → - P.sym $ pam-lemma xs g (pure ∘ f)) ⟩ + ≡.sym $ pam-lemma xs g (pure ∘ f)) ⟩ (fs >>= λ f → gs >>= λ g → pam (pam xs g) f) ≡⟨ MP.cong (refl {x = fs}) (λ f → MP.associative gs (pam xs) (pure ∘ f)) ⟩ (fs >>= pam (gs >>= pam xs)) ≡⟨ unfold-⊛ fs (gs >>= pam xs) ⟨ fs ⊛ (gs >>= pam xs) - ≡⟨ P.cong (fs ⊛_) (unfold-⊛ gs xs) ⟨ + ≡⟨ ≡.cong (fs ⊛_) (unfold-⊛ gs xs) ⟨ fs ⊛ (gs ⊛ xs) ∎ @@ -303,6 +303,6 @@ module Applicative where (fs >>= pam (pure x)) ≡⟨ (MP.cong (refl {x = fs}) λ f → MP.left-identity x (pure ∘ f)) ⟩ (fs >>= λ f → pure (f x)) ≡⟨⟩ - (pam fs (_$′ x)) ≡⟨ P.sym $ MP.left-identity (_$′ x) (pam fs) ⟩ + (pam fs (_$′ x)) ≡⟨ ≡.sym $ MP.left-identity (_$′ x) (pam fs) ⟩ (pure (_$′ x) >>= pam fs) ≡⟨ unfold-⊛ (pure (_$′ x)) fs ⟨ pure (_$′ x) ⊛ fs ∎ diff --git a/src/Data/List/Relation/Binary/Equality/Setoid.agda b/src/Data/List/Relation/Binary/Equality/Setoid.agda index fc6842c38d..af78f43c2c 100644 --- a/src/Data/List/Relation/Binary/Equality/Setoid.agda +++ b/src/Data/List/Relation/Binary/Equality/Setoid.agda @@ -20,7 +20,7 @@ open import Data.List.Relation.Unary.Unique.Setoid S using (Unique) open import Function.Base using (_∘_) open import Level open import Relation.Binary.Core renaming (Rel to Rel₂) -open import Relation.Binary.PropositionalEquality as P using (_≡_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) open import Relation.Binary.Properties.Setoid S using (≉-resp₂) open import Relation.Unary as U using (Pred) @@ -50,7 +50,7 @@ open PW public ≋-refl = PW.refl refl ≋-reflexive : _≡_ ⇒ _≋_ -≋-reflexive P.refl = ≋-refl +≋-reflexive ≡.refl = ≋-refl ≋-sym : Symmetric _≋_ ≋-sym = PW.symmetric sym diff --git a/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Solver.agda b/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Solver.agda index 1813568d5b..86194916f8 100644 --- a/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Solver.agda +++ b/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Solver.agda @@ -32,11 +32,11 @@ open import Data.List.Relation.Binary.Sublist.Heterogeneous open import Data.List.Relation.Binary.Sublist.Heterogeneous.Properties open import Function -open import Relation.Binary.PropositionalEquality as P +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_; sym; cong; cong₂; subst₂) open import Relation.Nullary -open P.≡-Reasoning +open ≡.≡-Reasoning infix 4 _⊆I_ _⊆R_ _⊆T_ @@ -87,7 +87,7 @@ d ⊆R e = ∀ ρ → Sublist R (⟦ d ⟧R ρ) (⟦ e ⟧R ρ) -- Flattening in a semantics-respecting manner ⟦++⟧R : ∀ {n} xs ys (ρ : Vec (List A) n) → ⟦ xs ++ ys ⟧R ρ ≡ ⟦ xs ⟧R ρ ++ ⟦ ys ⟧R ρ -⟦++⟧R [] ys ρ = P.refl +⟦++⟧R [] ys ρ = ≡.refl ⟦++⟧R (x ∷ xs) ys ρ = begin ⟦ x ⟧I ρ ++ ⟦ xs ++ ys ⟧R ρ ≡⟨ cong (⟦ x ⟧I ρ ++_) (⟦++⟧R xs ys ρ) ⟩ @@ -97,7 +97,7 @@ d ⊆R e = ∀ ρ → Sublist R (⟦ d ⟧R ρ) (⟦ e ⟧R ρ) ∎ flatten : ∀ {n} (t : TList n) → Σ[ r ∈ RList n ] ⟦ r ⟧R ≗ ⟦ t ⟧T -flatten [] = [] , λ _ → P.refl +flatten [] = [] , λ _ → ≡.refl flatten (It it) = it ∷ [] , λ ρ → ++-identityʳ (⟦ It it ⟧T ρ) flatten (t <> u) = let (rt , eqt) = flatten t diff --git a/src/Data/List/Relation/Unary/First/Properties.agda b/src/Data/List/Relation/Unary/First/Properties.agda index f30e037022..d598f14469 100644 --- a/src/Data/List/Relation/Unary/First/Properties.agda +++ b/src/Data/List/Relation/Unary/First/Properties.agda @@ -16,7 +16,7 @@ open import Data.List.Relation.Unary.Any as Any using (here; there) open import Data.List.Relation.Unary.First import Data.Sum as Sum open import Function.Base using (_∘′_; _$_; _∘_; id) -open import Relation.Binary.PropositionalEquality as P using (_≡_; refl; _≗_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl; _≗_) open import Relation.Unary open import Relation.Nullary.Negation @@ -66,14 +66,14 @@ module _ {a p q} {A : Set a} {P : Pred A p} {Q : Pred A q} where unique-index p⇒¬q [ _ ] [ _ ] = refl unique-index p⇒¬q [ qx ] (px ∷ _) = ⊥-elim (p⇒¬q px qx) unique-index p⇒¬q (px ∷ _) [ qx ] = ⊥-elim (p⇒¬q px qx) - unique-index p⇒¬q (_ ∷ f₁) (_ ∷ f₂) = P.cong suc (unique-index p⇒¬q f₁ f₂) + unique-index p⇒¬q (_ ∷ f₁) (_ ∷ f₂) = ≡.cong suc (unique-index p⇒¬q f₁ f₂) irrelevant : P ⊆ ∁ Q → Irrelevant P → Irrelevant Q → Irrelevant (First P Q) - irrelevant p⇒¬q p-irr q-irr [ qx₁ ] [ qx₂ ] = P.cong [_] (q-irr qx₁ qx₂) + irrelevant p⇒¬q p-irr q-irr [ qx₁ ] [ qx₂ ] = ≡.cong [_] (q-irr qx₁ qx₂) irrelevant p⇒¬q p-irr q-irr [ qx₁ ] (px₂ ∷ f₂) = ⊥-elim (p⇒¬q px₂ qx₁) irrelevant p⇒¬q p-irr q-irr (px₁ ∷ f₁) [ qx₂ ] = ⊥-elim (p⇒¬q px₁ qx₂) irrelevant p⇒¬q p-irr q-irr (px₁ ∷ f₁) (px₂ ∷ f₂) = - P.cong₂ _∷_ (p-irr px₁ px₂) (irrelevant p⇒¬q p-irr q-irr f₁ f₂) + ≡.cong₂ _∷_ (p-irr px₁ px₂) (irrelevant p⇒¬q p-irr q-irr f₁ f₂) ------------------------------------------------------------------------ -- Decidability @@ -97,11 +97,11 @@ module _ {a p} {A : Set a} {P : Pred A p} where fromAny∘toAny≗id : ∀ {xs} → fromAny {Q = P} {x = xs} ∘′ toAny ≗ id fromAny∘toAny≗id [ qx ] = refl - fromAny∘toAny≗id (px ∷ pqxs) = P.cong (_ ∷_) (fromAny∘toAny≗id pqxs) + fromAny∘toAny≗id (px ∷ pqxs) = ≡.cong (_ ∷_) (fromAny∘toAny≗id pqxs) toAny∘fromAny≗id : ∀ {xs} → toAny {Q = P} ∘′ fromAny {x = xs} ≗ id toAny∘fromAny≗id (here px) = refl - toAny∘fromAny≗id (there v) = P.cong there (toAny∘fromAny≗id v) + toAny∘fromAny≗id (there v) = ≡.cong there (toAny∘fromAny≗id v) ------------------------------------------------------------------------ -- Equivalence between the inductive definition and the view diff --git a/src/Data/Product/Function/Dependent/Propositional.agda b/src/Data/Product/Function/Dependent/Propositional.agda index b99e1934f8..2063e1cfbb 100644 --- a/src/Data/Product/Function/Dependent/Propositional.agda +++ b/src/Data/Product/Function/Dependent/Propositional.agda @@ -24,7 +24,7 @@ open import Function.Consequences.Propositional using (inverseʳ⇒injective; strictlySurjective⇒surjective) open import Function.Definitions using (Inverseˡ; Inverseʳ; Injective; StrictlySurjective) open import Function.Bundles -open import Relation.Binary.PropositionalEquality as P using (_≡_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) private variable @@ -54,7 +54,7 @@ module _ where Σ I A ⇔ Σ J B Σ-⇔ {B = B} I↠J A⇔B = mk⇔ (map (to I↠J) (Equivalence.to A⇔B)) - (map (to⁻ I↠J) (Equivalence.from A⇔B ∘ P.subst B (P.sym (proj₂ (surjective I↠J _) P.refl)))) + (map (to⁻ I↠J) (Equivalence.from A⇔B ∘ ≡.subst B (≡.sym (proj₂ (surjective I↠J _) ≡.refl)))) -- See also Data.Product.Relation.Binary.Pointwise.Dependent.WithK.↣. @@ -68,7 +68,7 @@ module _ where Σ I A ↣ Σ J B Σ-↣ {I = I} {J = J} {A = A} {B = B} I↔J A↣B = mk↣ to-injective where - open P.≡-Reasoning + open ≡.≡-Reasoning I≃J = ↔⇒≃ I↔J @@ -77,42 +77,42 @@ module _ where {x₁ x₂ : I} {y : A (from (to x₁))} (g : ∀ x → A (from (to x)) → B (to x)) (eq : to x₁ ≡ to x₂) → - P.subst B eq (g x₁ y) ≡ g x₂ (P.subst A (P.cong from eq) y) + ≡.subst B eq (g x₁ y) ≡ g x₂ (≡.subst A (≡.cong from eq) y) subst-application′ {x₁} {x₂} {y} g eq = - P.subst B eq (g x₁ y) ≡⟨ P.cong (P.subst B eq) (P.sym (g′-lemma _ _)) ⟩ - P.subst B eq (g′ (to x₁) y) ≡⟨ P.subst-application A g′ eq ⟩ - g′ (to x₂) (P.subst A (P.cong from eq) y) ≡⟨ g′-lemma _ _ ⟩ - g x₂ (P.subst A (P.cong from eq) y) ∎ + ≡.subst B eq (g x₁ y) ≡⟨ ≡.cong (≡.subst B eq) (≡.sym (g′-lemma _ _)) ⟩ + ≡.subst B eq (g′ (to x₁) y) ≡⟨ ≡.subst-application A g′ eq ⟩ + g′ (to x₂) (≡.subst A (≡.cong from eq) y) ≡⟨ g′-lemma _ _ ⟩ + g x₂ (≡.subst A (≡.cong from eq) y) ∎ where open _≃_ I≃J g′ : ∀ x → A (from x) → B x g′ x = - P.subst B (right-inverse-of x) ∘ + ≡.subst B (right-inverse-of x) ∘ g (from x) ∘ - P.subst A (P.sym (P.cong from (right-inverse-of x))) + ≡.subst A (≡.sym (≡.cong from (right-inverse-of x))) g′-lemma : ∀ x y → g′ (to x) y ≡ g x y g′-lemma x y = - P.subst B (right-inverse-of (to x)) + ≡.subst B (right-inverse-of (to x)) (g (from (to x)) $ - P.subst A (P.sym (P.cong from (right-inverse-of (to x)))) y) ≡⟨ P.cong (λ p → P.subst B p (g (from (to x)) - (P.subst A (P.sym (P.cong from p)) y))) - (P.sym (left-right x)) ⟩ - P.subst B (P.cong to (left-inverse-of x)) + ≡.subst A (≡.sym (≡.cong from (right-inverse-of (to x)))) y) ≡⟨ ≡.cong (λ p → ≡.subst B p (g (from (to x)) + (≡.subst A (≡.sym (≡.cong from p)) y))) + (≡.sym (left-right x)) ⟩ + ≡.subst B (≡.cong to (left-inverse-of x)) (g (from (to x)) $ - P.subst A - (P.sym (P.cong from (P.cong to (left-inverse-of x)))) + ≡.subst A + (≡.sym (≡.cong from (≡.cong to (left-inverse-of x)))) y) ≡⟨ lemma _ ⟩ g x y ∎ where lemma : ∀ {x′} eq {y : A (from (to x′))} → - P.subst B (P.cong to eq) + ≡.subst B (≡.cong to eq) (g (from (to x)) - (P.subst A (P.sym (P.cong from (P.cong to eq))) y)) ≡ + (≡.subst A (≡.sym (≡.cong from (≡.cong to eq))) y)) ≡ g x′ y - lemma P.refl = P.refl + lemma ≡.refl = ≡.refl open Injection @@ -125,49 +125,49 @@ module _ where Σ-≡,≡→≡ ∘′ map (_≃_.injective I≃J) (λ {eq₁} eq₂ → injective A↣B ( - to A↣B (P.subst A (_≃_.injective I≃J eq₁) x₂) ≡⟨⟩ + to A↣B (≡.subst A (_≃_.injective I≃J eq₁) x₂) ≡⟨⟩ (let eq = - P.trans (P.sym (_≃_.left-inverse-of I≃J x₁)) - (P.trans (P.cong (_≃_.from I≃J) eq₁) - (P.trans (_≃_.left-inverse-of I≃J y₁) - P.refl)) in - to A↣B (P.subst A eq x₂)) ≡⟨ P.cong (λ p → to A↣B - (P.subst A - (P.trans (P.sym (_≃_.left-inverse-of I≃J _)) - (P.trans (P.cong (_≃_.from I≃J) eq₁) p)) + ≡.trans (≡.sym (_≃_.left-inverse-of I≃J x₁)) + (≡.trans (≡.cong (_≃_.from I≃J) eq₁) + (≡.trans (_≃_.left-inverse-of I≃J y₁) + ≡.refl)) in + to A↣B (≡.subst A eq x₂)) ≡⟨ ≡.cong (λ p → to A↣B + (≡.subst A + (≡.trans (≡.sym (_≃_.left-inverse-of I≃J _)) + (≡.trans (≡.cong (_≃_.from I≃J) eq₁) p)) x₂)) - (P.trans-reflʳ _) ⟩ + (≡.trans-reflʳ _) ⟩ - (let eq = P.trans (P.sym (_≃_.left-inverse-of I≃J x₁)) - (P.trans (P.cong (_≃_.from I≃J) eq₁) + (let eq = ≡.trans (≡.sym (_≃_.left-inverse-of I≃J x₁)) + (≡.trans (≡.cong (_≃_.from I≃J) eq₁) (_≃_.left-inverse-of I≃J y₁)) in - to A↣B (P.subst A eq x₂)) ≡⟨ P.cong (to A↣B) - (P.sym (P.subst-subst (P.sym (_≃_.left-inverse-of I≃J _)))) ⟩ + to A↣B (≡.subst A eq x₂)) ≡⟨ ≡.cong (to A↣B) + (≡.sym (≡.subst-subst (≡.sym (_≃_.left-inverse-of I≃J _)))) ⟩ - to A↣B ((P.subst A (P.trans (P.cong (_≃_.from I≃J) eq₁) + to A↣B ((≡.subst A (≡.trans (≡.cong (_≃_.from I≃J) eq₁) (_≃_.left-inverse-of I≃J y₁)) $ - P.subst A (P.sym (_≃_.left-inverse-of I≃J x₁)) x₂)) ≡⟨ P.cong (to A↣B) - (P.sym (P.subst-subst (P.cong (_≃_.from I≃J) eq₁))) ⟩ + ≡.subst A (≡.sym (_≃_.left-inverse-of I≃J x₁)) x₂)) ≡⟨ ≡.cong (to A↣B) + (≡.sym (≡.subst-subst (≡.cong (_≃_.from I≃J) eq₁))) ⟩ to A↣B ( - (P.subst A (_≃_.left-inverse-of I≃J y₁) $ - P.subst A (P.cong (_≃_.from I≃J) eq₁) $ - P.subst A (P.sym (_≃_.left-inverse-of I≃J x₁)) x₂)) ≡⟨ P.sym (subst-application′ + (≡.subst A (_≃_.left-inverse-of I≃J y₁) $ + ≡.subst A (≡.cong (_≃_.from I≃J) eq₁) $ + ≡.subst A (≡.sym (_≃_.left-inverse-of I≃J x₁)) x₂)) ≡⟨ ≡.sym (subst-application′ (λ x y → to A↣B - (P.subst A (_≃_.left-inverse-of I≃J x) y)) + (≡.subst A (_≃_.left-inverse-of I≃J x) y)) eq₁) ⟩ - P.subst B eq₁ (to A↣B $ - (P.subst A (_≃_.left-inverse-of I≃J x₁) $ - P.subst A (P.sym (_≃_.left-inverse-of I≃J x₁)) x₂)) ≡⟨ P.cong (P.subst B eq₁ ∘ to A↣B) - (P.subst-subst (P.sym (_≃_.left-inverse-of I≃J _))) ⟩ + ≡.subst B eq₁ (to A↣B $ + (≡.subst A (_≃_.left-inverse-of I≃J x₁) $ + ≡.subst A (≡.sym (_≃_.left-inverse-of I≃J x₁)) x₂)) ≡⟨ ≡.cong (≡.subst B eq₁ ∘ to A↣B) + (≡.subst-subst (≡.sym (_≃_.left-inverse-of I≃J _))) ⟩ - (let eq = P.trans (P.sym (_≃_.left-inverse-of I≃J x₁)) + (let eq = ≡.trans (≡.sym (_≃_.left-inverse-of I≃J x₁)) (_≃_.left-inverse-of I≃J x₁) in - P.subst B eq₁ (to A↣B (P.subst A eq x₂))) ≡⟨ P.cong (λ p → P.subst B eq₁ (to A↣B (P.subst A p x₂))) - (P.trans-symˡ (_≃_.left-inverse-of I≃J _)) ⟩ - P.subst B eq₁ (to A↣B (P.subst A P.refl x₂)) ≡⟨⟩ + ≡.subst B eq₁ (to A↣B (≡.subst A eq x₂))) ≡⟨ ≡.cong (λ p → ≡.subst B eq₁ (to A↣B (≡.subst A p x₂))) + (≡.trans-symˡ (_≃_.left-inverse-of I≃J _)) ⟩ + ≡.subst B eq₁ (to A↣B (≡.subst A ≡.refl x₂)) ≡⟨⟩ - P.subst B eq₁ (to A↣B x₂) ≡⟨ eq₂ ⟩ + ≡.subst B eq₁ (to A↣B x₂) ≡⟨ eq₂ ⟩ to A↣B y₂ ∎ @@ -191,7 +191,7 @@ module _ where to′ = map (to I↠J) (to A↠B) backcast : ∀ {i} → B i → B (to I↠J (to⁻ I↠J i)) - backcast = P.subst B (P.sym (to∘to⁻ I↠J _)) + backcast = ≡.subst B (≡.sym (to∘to⁻ I↠J _)) to⁻′ : Σ J B → Σ I A to⁻′ = map (to⁻ I↠J) (Surjection.to⁻ A↠B ∘ backcast) @@ -199,10 +199,10 @@ module _ where strictlySurjective′ : StrictlySurjective _≡_ to′ strictlySurjective′ (x , y) = to⁻′ (x , y) , Σ-≡,≡→≡ ( to∘to⁻ I↠J x - , (P.subst B (to∘to⁻ I↠J x) (to A↠B (to⁻ A↠B (backcast y))) ≡⟨ P.cong (P.subst B _) (to∘to⁻ A↠B _) ⟩ - P.subst B (to∘to⁻ I↠J x) (backcast y) ≡⟨ P.subst-subst-sym (to∘to⁻ I↠J x) ⟩ + , (≡.subst B (to∘to⁻ I↠J x) (to A↠B (to⁻ A↠B (backcast y))) ≡⟨ ≡.cong (≡.subst B _) (to∘to⁻ A↠B _) ⟩ + ≡.subst B (to∘to⁻ I↠J x) (backcast y) ≡⟨ ≡.subst-subst-sym (to∘to⁻ I↠J x) ⟩ y ∎) - ) where open P.≡-Reasoning + ) where open ≡.≡-Reasoning ------------------------------------------------------------------------ @@ -220,17 +220,17 @@ module _ where to′ = map (to I↩J) (to A↩B) backcast : ∀ {j} → B j → B (to I↩J (from I↩J j)) - backcast = P.subst B (P.sym (inverseˡ I↩J P.refl)) + backcast = ≡.subst B (≡.sym (inverseˡ I↩J ≡.refl)) from′ : Σ J B → Σ I A from′ = map (from I↩J) (from A↩B ∘ backcast) inv : Inverseˡ _≡_ _≡_ to′ from′ - inv {j , b} P.refl = Σ-≡,≡→≡ (strictlyInverseˡ I↩J j , ( + inv {j , b} ≡.refl = Σ-≡,≡→≡ (strictlyInverseˡ I↩J j , ( begin - P.subst B (inverseˡ I↩J P.refl) (to A↩B (from A↩B (backcast b))) ≡⟨ P.cong (P.subst B _) (inverseˡ A↩B P.refl) ⟩ - P.subst B (inverseˡ I↩J P.refl) (backcast b) ≡⟨ P.subst-subst-sym (inverseˡ I↩J _) ⟩ - b ∎)) where open P.≡-Reasoning + ≡.subst B (inverseˡ I↩J ≡.refl) (to A↩B (from A↩B (backcast b))) ≡⟨ ≡.cong (≡.subst B _) (inverseˡ A↩B ≡.refl) ⟩ + ≡.subst B (inverseˡ I↩J ≡.refl) (backcast b) ≡⟨ ≡.subst-subst-sym (inverseˡ I↩J _) ⟩ + b ∎)) where open ≡.≡-Reasoning ------------------------------------------------------------------------ -- Right inverses @@ -250,7 +250,7 @@ module _ where (Surjection.to∘to⁻ surjection′) left-inverse-of where - open P.≡-Reasoning + open ≡.≡-Reasoning I≃J = ↔⇒≃ I↔J @@ -260,27 +260,27 @@ module _ where left-inverse-of : ∀ p → Surjection.to⁻ surjection′ (Surjection.to surjection′ p) ≡ p left-inverse-of (x , y) = to Σ-≡,≡↔≡ ( _≃_.left-inverse-of I≃J x - , (P.subst A (_≃_.left-inverse-of I≃J x) + , (≡.subst A (_≃_.left-inverse-of I≃J x) (from A↔B - (P.subst B (P.sym (_≃_.right-inverse-of I≃J + (≡.subst B (≡.sym (_≃_.right-inverse-of I≃J (_≃_.to I≃J x))) - (to A↔B y))) ≡⟨ P.subst-application B (λ _ → from A↔B) _ ⟩ + (to A↔B y))) ≡⟨ ≡.subst-application B (λ _ → from A↔B) _ ⟩ from A↔B - (P.subst B (P.cong (_≃_.to I≃J) + (≡.subst B (≡.cong (_≃_.to I≃J) (_≃_.left-inverse-of I≃J x)) - (P.subst B (P.sym (_≃_.right-inverse-of I≃J + (≡.subst B (≡.sym (_≃_.right-inverse-of I≃J (_≃_.to I≃J x))) - (to A↔B y))) ≡⟨ P.cong (λ eq → from A↔B (P.subst B eq - (P.subst B (P.sym (_≃_.right-inverse-of I≃J _)) _))) + (to A↔B y))) ≡⟨ ≡.cong (λ eq → from A↔B (≡.subst B eq + (≡.subst B (≡.sym (_≃_.right-inverse-of I≃J _)) _))) (_≃_.left-right I≃J _) ⟩ from A↔B - (P.subst B (_≃_.right-inverse-of I≃J + (≡.subst B (_≃_.right-inverse-of I≃J (_≃_.to I≃J x)) - (P.subst B (P.sym (_≃_.right-inverse-of I≃J + (≡.subst B (≡.sym (_≃_.right-inverse-of I≃J (_≃_.to I≃J x))) - (to A↔B y))) ≡⟨ P.cong (from A↔B) - (P.subst-subst-sym (_≃_.right-inverse-of I≃J _)) ⟩ + (to A↔B y))) ≡⟨ ≡.cong (from A↔B) + (≡.subst-subst-sym (_≃_.right-inverse-of I≃J _)) ⟩ from A↔B (to A↔B y) ≡⟨ Inverse.strictlyInverseʳ A↔B _ ⟩ @@ -297,7 +297,7 @@ private module _ where ∀ {x} → A (from I↔J x) ∼[ k ] B x swap-coercions {A = A} B I↔J eq {x} = A (from I↔J x) ∼⟨ eq ⟩ - B (to I↔J (from I↔J x)) ↔⟨ K-reflexive (P.cong B $ strictlyInverseˡ I↔J x) ⟩ + B (to I↔J (from I↔J x)) ↔⟨ K-reflexive (≡.cong B $ strictlyInverseˡ I↔J x) ⟩ B x ∎ where open EquationalReasoning diff --git a/src/Data/Product/Function/Dependent/Propositional/WithK.agda b/src/Data/Product/Function/Dependent/Propositional/WithK.agda index 83aa4a9baa..920bb2c136 100644 --- a/src/Data/Product/Function/Dependent/Propositional/WithK.agda +++ b/src/Data/Product/Function/Dependent/Propositional/WithK.agda @@ -20,7 +20,7 @@ open import Level using (Level) open import Function open import Function.Properties.Injection open import Function.Properties.Inverse as Inverse -open import Relation.Binary.PropositionalEquality as P using (_≡_; refl) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl) private variable diff --git a/src/Function/Bijection.agda b/src/Function/Bijection.agda index ac32159e7a..91974b6294 100644 --- a/src/Function/Bijection.agda +++ b/src/Function/Bijection.agda @@ -16,7 +16,7 @@ Use the standard function hierarchy in Function/Function.Bundles instead." open import Level open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality as P +open import Relation.Binary.PropositionalEquality as ≡ open import Function.Equality as F using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_) open import Function.Injection as Inj hiding (id; _∘_; injection) @@ -90,7 +90,7 @@ Please use Function.(Bundles.)Bijection instead." infix 3 _⤖_ _⤖_ : ∀ {f t} → Set f → Set t → Set _ -From ⤖ To = Bijection (P.setoid From) (P.setoid To) +From ⤖ To = Bijection (≡.setoid From) (≡.setoid To) {-# WARNING_ON_USAGE _⤖_ "Warning: _⤖_ was deprecated in v2.0. Please use Function.(Bundles.)mk⤖ instead." diff --git a/src/Function/Equivalence.agda b/src/Function/Equivalence.agda index 1e3652da7f..9556023b56 100644 --- a/src/Function/Equivalence.agda +++ b/src/Function/Equivalence.agda @@ -20,7 +20,7 @@ open import Function.Equality as F open import Level open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (Reflexive; TransFlip; Sym) -import Relation.Binary.PropositionalEquality as P +import Relation.Binary.PropositionalEquality as ≡ ------------------------------------------------------------------------ -- Setoid equivalence @@ -43,7 +43,7 @@ Please use Function.(Bundles.)Equivalence instead." infix 3 _⇔_ _⇔_ : ∀ {f t} → Set f → Set t → Set _ -From ⇔ To = Equivalence (P.setoid From) (P.setoid To) +From ⇔ To = Equivalence (≡.setoid From) (≡.setoid To) {-# WARNING_ON_USAGE _⇔_ "Warning: _⇔_ was deprecated in v2.0. Please use Function.(Bundles.)_⇔_ instead." diff --git a/src/Function/Injection.agda b/src/Function/Injection.agda index 462b8b3428..72785584f1 100644 --- a/src/Function/Injection.agda +++ b/src/Function/Injection.agda @@ -19,7 +19,7 @@ open import Level open import Relation.Binary.Bundles using (Setoid) open import Function.Equality as F using (_⟶_; _⟨$⟩_ ; Π) renaming (_∘_ to _⟪∘⟫_) -open import Relation.Binary.PropositionalEquality as P using (_≡_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) ------------------------------------------------------------------------ -- Injective functions @@ -58,7 +58,7 @@ Please use Function.(Bundles.)Injection instead." infix 3 _↣_ _↣_ : ∀ {f t} → Set f → Set t → Set _ -From ↣ To = Injection (P.setoid From) (P.setoid To) +From ↣ To = Injection (≡.setoid From) (≡.setoid To) {-# WARNING_ON_USAGE _↣_ "Warning: _↣_ was deprecated in v2.0. Please use Function.(Bundles.)_↣_ instead." @@ -69,7 +69,7 @@ injection : ∀ {f t} {From : Set f} {To : Set t} → (to : From → To) → injection to injective = record { to = record { _⟨$⟩_ = to - ; cong = P.cong to + ; cong = ≡.cong to } ; injective = injective } diff --git a/src/Function/Inverse.agda b/src/Function/Inverse.agda index e9d796eddf..792382ca17 100644 --- a/src/Function/Inverse.agda +++ b/src/Function/Inverse.agda @@ -22,7 +22,7 @@ open import Function.Equality as F open import Function.LeftInverse as Left hiding (id; _∘_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (Reflexive; TransFlip; Sym) -open import Relation.Binary.PropositionalEquality as P using (_≗_; _≡_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_; _≡_) open import Relation.Unary using (Pred) ------------------------------------------------------------------------ @@ -90,7 +90,7 @@ Please use Function.(Bundles.)Inverse instead." infix 3 _↔_ _↔̇_ _↔_ : ∀ {f t} → Set f → Set t → Set _ -From ↔ To = Inverse (P.setoid From) (P.setoid To) +From ↔ To = Inverse (≡.setoid From) (≡.setoid To) {-# WARNING_ON_USAGE _↔_ "Warning: _↔_ was deprecated in v2.0. Please use Function.(Bundles.)_↔_ instead." diff --git a/src/Function/LeftInverse.agda b/src/Function/LeftInverse.agda index 3824ca95d3..d9ccc9a806 100644 --- a/src/Function/LeftInverse.agda +++ b/src/Function/LeftInverse.agda @@ -21,7 +21,7 @@ open import Function.Equality as Eq using (_⟶_; _⟨$⟩_) renaming (_∘_ to _⟪∘⟫_) open import Function.Equivalence using (Equivalence) open import Function.Injection using (Injective; Injection) -open import Relation.Binary.PropositionalEquality as P using (_≡_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) ------------------------------------------------------------------------ -- Left and right inverses. @@ -106,7 +106,7 @@ Please use Function.(Bundles.)LeftInverse instead." infix 3 _↞_ _↞_ : ∀ {f t} → Set f → Set t → Set _ -From ↞ To = LeftInverse (P.setoid From) (P.setoid To) +From ↞ To = LeftInverse (≡.setoid From) (≡.setoid To) {-# WARNING_ON_USAGE _↞_ "Warning: _↞_ was deprecated in v2.0. Please use Function.(Bundles.)_↪_ instead." diff --git a/src/Function/Surjection.agda b/src/Function/Surjection.agda index 2fdfefa9af..684da0b3d7 100644 --- a/src/Function/Surjection.agda +++ b/src/Function/Surjection.agda @@ -21,7 +21,7 @@ open import Function.Equivalence using (Equivalence) open import Function.Injection hiding (id; _∘_; injection) open import Function.LeftInverse as Left hiding (id; _∘_) open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality as P using (_≡_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) ------------------------------------------------------------------------ -- Surjective functions. @@ -100,7 +100,7 @@ Please use Function.(Properties.)RightInverse.RightInverse⇒Surjection instead. infix 3 _↠_ _↠_ : ∀ {f t} → Set f → Set t → Set _ -From ↠ To = Surjection (P.setoid From) (P.setoid To) +From ↠ To = Surjection (≡.setoid From) (≡.setoid To) {-# WARNING_ON_USAGE _↠_ "Warning: _↠_ was deprecated in v2.0. Please use Function.(Bundles.)_↠_ instead." diff --git a/src/Relation/Binary/Structures.agda b/src/Relation/Binary/Structures.agda index 0dc5187228..14426b4cff 100644 --- a/src/Relation/Binary/Structures.agda +++ b/src/Relation/Binary/Structures.agda @@ -18,7 +18,7 @@ module Relation.Binary.Structures open import Data.Product.Base using (proj₁; proj₂; _,_) open import Level using (Level; _⊔_) open import Relation.Nullary.Negation.Core using (¬_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Consequences open import Relation.Binary.Definitions @@ -51,7 +51,7 @@ record IsEquivalence : Set (a ⊔ ℓ) where trans : Transitive _≈_ reflexive : _≡_ ⇒ _≈_ - reflexive P.refl = refl + reflexive ≡.refl = refl isPartialEquivalence : IsPartialEquivalence isPartialEquivalence = record From 46ce8c8b5a4fcc8da9d225b2a78da5c9bf14efc0 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 13 Feb 2024 07:59:27 +0000 Subject: [PATCH 2/5] Qualified import of `PropositionalEquality.Core` fixing #2280 --- src/Algebra/Solver/Monoid.agda | 4 +- src/Algebra/Solver/Ring.agda | 4 +- .../Ring/NaturalCoefficients/Default.agda | 4 +- src/Codata/Guarded/Stream/Properties.agda | 84 +++++++++---------- src/Codata/Musical/Colist/Infinite-merge.agda | 16 ++-- src/Codata/Musical/Conat.agda | 10 +-- src/Codata/Musical/Covec.agda | 6 +- src/Codata/Musical/Stream.agda | 12 +-- src/Codata/Sized/Colist/Bisimilarity.agda | 14 ++-- src/Codata/Sized/Covec/Bisimilarity.agda | 8 +- src/Codata/Sized/Covec/Properties.agda | 6 +- src/Codata/Sized/Cowriter/Bisimilarity.agda | 14 ++-- src/Codata/Sized/Delay/Bisimilarity.agda | 8 +- src/Codata/Sized/Delay/Properties.agda | 20 ++--- src/Codata/Sized/M/Bisimilarity.agda | 10 +-- src/Codata/Sized/M/Properties.agda | 14 ++-- src/Codata/Sized/Stream/Bisimilarity.agda | 12 +-- src/Codata/Sized/Stream/Properties.agda | 46 +++++----- .../Binary/Equality/Propositional.agda | 12 +-- src/Effect/Applicative.agda | 2 +- src/Effect/Applicative/Indexed.agda | 4 +- src/Effect/Monad/Partiality.agda | 42 +++++----- src/Effect/Monad/Partiality/All.agda | 10 +-- src/Function/Endomorphism/Propositional.agda | 16 ++-- src/Function/Equality.agda | 8 +- src/Function/Related.agda | 30 +++---- src/Function/Related/Propositional.agda | 16 ++-- src/Function/Related/TypeIsomorphisms.agda | 48 +++++------ .../Construct/Add/Infimum/NonStrict.agda | 12 +-- .../Binary/Construct/Add/Infimum/Strict.agda | 14 ++-- .../Binary/Construct/Add/Point/Equality.agda | 6 +- .../Construct/Add/Supremum/NonStrict.agda | 12 +-- .../Binary/Construct/Add/Supremum/Strict.agda | 14 ++-- .../Closure/Reflexive/Properties.agda | 6 +- .../Binary/Construct/Closure/Transitive.agda | 4 +- .../Binary/HeterogeneousEquality.agda | 16 ++-- .../Binary/Indexed/Heterogeneous/Bundles.agda | 2 +- .../Binary/Indexed/Heterogeneous/Core.agda | 2 +- .../Indexed/Heterogeneous/Definitions.agda | 2 +- .../Indexed/Heterogeneous/Structures.agda | 4 +- .../Binary/Indexed/Homogeneous/Bundles.agda | 2 +- .../Indexed/Homogeneous/Structures.agda | 6 +- src/Relation/Binary/Properties/Setoid.agda | 8 +- src/Tactic/Cong.agda | 2 +- 44 files changed, 296 insertions(+), 296 deletions(-) diff --git a/src/Algebra/Solver/Monoid.agda b/src/Algebra/Solver/Monoid.agda index f80a2847f6..2aea92797c 100644 --- a/src/Algebra/Solver/Monoid.agda +++ b/src/Algebra/Solver/Monoid.agda @@ -22,7 +22,7 @@ open import Data.Vec.Base using (Vec; lookup) open import Function.Base using (_∘_; _$_) open import Relation.Binary.Definitions using (Decidable) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) import Relation.Binary.Reflection open import Relation.Nullary import Relation.Nullary.Decidable as Dec @@ -128,7 +128,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/Ring.agda b/src/Algebra/Solver/Ring.agda index 50c191903b..a2a982d5c2 100644 --- a/src/Algebra/Solver/Ring.agda +++ b/src/Algebra/Solver/Ring.agda @@ -41,7 +41,7 @@ open import Algebra.Properties.Semiring.Exp semiring open import Relation.Nullary.Decidable using (yes; no) open import Relation.Binary.Reasoning.Setoid setoid -import Relation.Binary.PropositionalEquality.Core as PropEq +import Relation.Binary.PropositionalEquality.Core as ≡ import Relation.Binary.Reflection as Reflection open import Data.Nat.Base using (ℕ; suc; zero) @@ -534,7 +534,7 @@ correct (con c) ρ = correct-con c ρ correct (var i) ρ = correct-var i ρ correct (p :^ k) ρ = begin ⟦ normalise p ^N k ⟧N ρ ≈⟨ ^N-homo (normalise p) k ρ ⟩ - ⟦ p ⟧↓ ρ ^ k ≈⟨ correct p ρ ⟨ ^-cong ⟩ PropEq.refl {x = k} ⟩ + ⟦ p ⟧↓ ρ ^ k ≈⟨ correct p ρ ⟨ ^-cong ⟩ ≡.refl {x = k} ⟩ ⟦ p ⟧ ρ ^ k ∎ correct (:- p) ρ = begin ⟦ -N normalise p ⟧N ρ ≈⟨ -N‿-homo (normalise p) ρ ⟩ diff --git a/src/Algebra/Solver/Ring/NaturalCoefficients/Default.agda b/src/Algebra/Solver/Ring/NaturalCoefficients/Default.agda index 9b9671c527..da1b56b93f 100644 --- a/src/Algebra/Solver/Ring/NaturalCoefficients/Default.agda +++ b/src/Algebra/Solver/Ring/NaturalCoefficients/Default.agda @@ -20,13 +20,13 @@ import Algebra.Properties.Semiring.Mult as SemiringMultiplication open import Data.Maybe.Base using (Maybe; map) open import Data.Nat using (_≟_) open import Relation.Binary.Consequences using (dec⇒weaklyDec) -import Relation.Binary.PropositionalEquality.Core as P +import Relation.Binary.PropositionalEquality.Core as ≡ open CommutativeSemiring R open SemiringMultiplication semiring private dec : ∀ m n → Maybe (m × 1# ≈ n × 1#) - dec m n = map (λ { P.refl → refl }) (dec⇒weaklyDec _≟_ m n) + dec m n = map (λ { ≡.refl → refl }) (dec⇒weaklyDec _≟_ m n) open import Algebra.Solver.Ring.NaturalCoefficients R dec public diff --git a/src/Codata/Guarded/Stream/Properties.agda b/src/Codata/Guarded/Stream/Properties.agda index 3742aab856..7a547e6ec2 100644 --- a/src/Codata/Guarded/Stream/Properties.agda +++ b/src/Codata/Guarded/Stream/Properties.agda @@ -20,7 +20,7 @@ open import Data.Product.Base as Prod using (_×_; _,_; proj₁; proj₂) open import Data.Vec.Base as Vec using (Vec; _∷_) open import Function.Base using (const; flip; id; _∘′_; _$′_; _⟨_⟩_; _∘₂′_) open import Level using (Level) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; cong; cong₂) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; cong; cong₂) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) @@ -39,7 +39,7 @@ cong-lookup : ∀ {as bs : Stream A} → as ≈ bs → ∀ n → lookup as n ≡ cong-lookup = B.lookup⁺ cong-take : ∀ n {as bs : Stream A} → as ≈ bs → take n as ≡ take n bs -cong-take zero as≈bs = P.refl +cong-take zero as≈bs = ≡.refl cong-take (suc n) as≈bs = cong₂ _∷_ (as≈bs .head) (cong-take n (as≈bs .tail)) cong-drop : ∀ n {as bs : Stream A} → as ≈ bs → drop n as ≈ drop n bs @@ -63,7 +63,7 @@ cong-concat ass≈bss = cong-++-concat [] ass≈bss cong-++-concat : ∀ (as : List A) {ass bss} → ass ≈ bss → ++-concat as ass ≈ ++-concat as bss cong-++-concat [] ass≈bss .head = cong List⁺.head (ass≈bss .head) cong-++-concat [] ass≈bss .tail rewrite ass≈bss .head = cong-++-concat _ (ass≈bss .tail) - cong-++-concat (a ∷ as) ass≈bss .head = P.refl + cong-++-concat (a ∷ as) ass≈bss .head = ≡.refl cong-++-concat (a ∷ as) ass≈bss .tail = cong-++-concat as ass≈bss cong-interleave : {as bs cs ds : Stream A} → as ≈ bs → cs ≈ ds → @@ -79,11 +79,11 @@ cong-chunksOf n as≈bs .tail = cong-chunksOf n (cong-drop n as≈bs) -- Properties of repeat lookup-repeat : ∀ n (a : A) → lookup (repeat a) n ≡ a -lookup-repeat zero a = P.refl +lookup-repeat zero a = ≡.refl lookup-repeat (suc n) a = lookup-repeat n a splitAt-repeat : ∀ n (a : A) → splitAt n (repeat a) ≡ (Vec.replicate n a , repeat a) -splitAt-repeat zero a = P.refl +splitAt-repeat zero a = ≡.refl splitAt-repeat (suc n) a = cong (Prod.map₁ (a ∷_)) (splitAt-repeat n a) take-repeat : ∀ n (a : A) → take n (repeat a) ≡ Vec.replicate n a @@ -93,28 +93,28 @@ drop-repeat : ∀ n (a : A) → drop n (repeat a) ≡ repeat a drop-repeat n a = cong proj₂ (splitAt-repeat n a) map-repeat : ∀ (f : A → B) a → map f (repeat a) ≈ repeat (f a) -map-repeat f a .head = P.refl +map-repeat f a .head = ≡.refl map-repeat f a .tail = map-repeat f a ap-repeat : ∀ (f : A → B) a → ap (repeat f) (repeat a) ≈ repeat (f a) -ap-repeat f a .head = P.refl +ap-repeat f a .head = ≡.refl ap-repeat f a .tail = ap-repeat f a ap-repeatˡ : ∀ (f : A → B) as → ap (repeat f) as ≈ map f as -ap-repeatˡ f as .head = P.refl +ap-repeatˡ f as .head = ≡.refl ap-repeatˡ f as .tail = ap-repeatˡ f (as .tail) ap-repeatʳ : ∀ (fs : Stream (A → B)) a → ap fs (repeat a) ≈ map (_$′ a) fs -ap-repeatʳ fs a .head = P.refl +ap-repeatʳ fs a .head = ≡.refl ap-repeatʳ fs a .tail = ap-repeatʳ (fs .tail) a interleave-repeat : (a : A) → interleave (repeat a) (repeat a) ≈ repeat a -interleave-repeat a .head = P.refl +interleave-repeat a .head = ≡.refl interleave-repeat a .tail = interleave-repeat a zipWith-repeat : ∀ (f : A → B → C) a b → zipWith f (repeat a) (repeat b) ≈ repeat (f a b) -zipWith-repeat f a b .head = P.refl +zipWith-repeat f a b .head = ≡.refl zipWith-repeat f a b .tail = zipWith-repeat f a b chunksOf-repeat : ∀ n (a : A) → chunksOf n (repeat a) ≈ repeat (Vec.replicate n a) @@ -125,7 +125,7 @@ chunksOf-repeat n a = begin go where go : chunksOf n (repeat a) ≈∞ repeat (Vec.replicate n a) go .head = take-repeat n a go .tail = - chunksOf n (drop n (repeat a)) ≡⟨ P.cong (chunksOf n) (drop-repeat n a) ⟩ + chunksOf n (drop n (repeat a)) ≡⟨ ≡.cong (chunksOf n) (drop-repeat n a) ⟩ chunksOf n (repeat a) ↺⟨ go ⟩ repeat (Vec.replicate n a) ∎ @@ -133,34 +133,34 @@ chunksOf-repeat n a = begin go where -- Properties of map map-const : (a : A) (bs : Stream B) → map (const a) bs ≈ repeat a -map-const a bs .head = P.refl +map-const a bs .head = ≡.refl map-const a bs .tail = map-const a (bs .tail) map-id : (as : Stream A) → map id as ≈ as -map-id as .head = P.refl +map-id as .head = ≡.refl map-id as .tail = map-id (as .tail) map-∘ : ∀ (g : B → C) (f : A → B) as → map g (map f as) ≈ map (g ∘′ f) as -map-∘ g f as .head = P.refl +map-∘ g f as .head = ≡.refl map-∘ g f as .tail = map-∘ g f (as .tail) map-unfold : ∀ (g : B → C) (f : A → A × B) a → map g (unfold f a) ≈ unfold (Prod.map₂ g ∘′ f) a -map-unfold g f a .head = P.refl +map-unfold g f a .head = ≡.refl map-unfold g f a .tail = map-unfold g f (proj₁ (f a)) map-drop : ∀ (f : A → B) n as → map f (drop n as) ≡ drop n (map f as) -map-drop f zero as = P.refl +map-drop f zero as = ≡.refl map-drop f (suc n) as = map-drop f n (as .tail) map-zipWith : ∀ (g : C → D) (f : A → B → C) as bs → map g (zipWith f as bs) ≈ zipWith (g ∘₂′ f) as bs -map-zipWith g f as bs .head = P.refl +map-zipWith g f as bs .head = ≡.refl map-zipWith g f as bs .tail = map-zipWith g f (as .tail) (bs .tail) map-interleave : ∀ (f : A → B) as bs → map f (interleave as bs) ≈ interleave (map f as) (map f bs) -map-interleave f as bs .head = P.refl +map-interleave f as bs .head = ≡.refl map-interleave f as bs .tail = map-interleave f bs (as .tail) map-concat : ∀ (f : A → B) ass → map f (concat ass) ≈ concat (map (List⁺.map f) ass) @@ -168,9 +168,9 @@ map-concat f ass = map-++-concat [] ass where open Concat map-++-concat : ∀ acc ass → map f (++-concat acc ass) ≈ ++-concat (List.map f acc) (map (List⁺.map f) ass) - map-++-concat [] ass .head = P.refl + map-++-concat [] ass .head = ≡.refl map-++-concat [] ass .tail = map-++-concat (ass .head .List⁺.tail) (ass .tail) - map-++-concat (a ∷ as) ass .head = P.refl + map-++-concat (a ∷ as) ass .head = ≡.refl map-++-concat (a ∷ as) ass .tail = map-++-concat as ass map-cycle : ∀ (f : A → B) as → map f (cycle as) ≈ cycle (List⁺.map f as) @@ -186,29 +186,29 @@ map-cycle f as = run -- Properties of lookup lookup-drop : ∀ m (as : Stream A) n → lookup (drop m as) n ≡ lookup as (m + n) -lookup-drop zero as n = P.refl +lookup-drop zero as n = ≡.refl lookup-drop (suc m) as n = lookup-drop m (as .tail) n lookup-map : ∀ n (f : A → B) as → lookup (map f as) n ≡ f (lookup as n) -lookup-map zero f as = P.refl +lookup-map zero f as = ≡.refl lookup-map (suc n) f as = lookup-map n f (as . tail) lookup-iterate : ∀ n f (x : A) → lookup (iterate f x) n ≡ ℕ.iterate f x n -lookup-iterate zero f x = P.refl +lookup-iterate zero f x = ≡.refl lookup-iterate (suc n) f x = lookup-iterate n f (f x) lookup-zipWith : ∀ n (f : A → B → C) as bs → lookup (zipWith f as bs) n ≡ f (lookup as n) (lookup bs n) -lookup-zipWith zero f as bs = P.refl +lookup-zipWith zero f as bs = ≡.refl lookup-zipWith (suc n) f as bs = lookup-zipWith n f (as .tail) (bs .tail) lookup-unfold : ∀ n (f : A → A × B) a → lookup (unfold f a) n ≡ proj₂ (f (ℕ.iterate (proj₁ ∘′ f) a n)) -lookup-unfold zero f a = P.refl +lookup-unfold zero f a = ≡.refl lookup-unfold (suc n) f a = lookup-unfold n f (proj₁ (f a)) lookup-tabulate : ∀ n (f : ℕ → A) → lookup (tabulate f) n ≡ f n -lookup-tabulate zero f = P.refl +lookup-tabulate zero f = ≡.refl lookup-tabulate (suc n) f = lookup-tabulate n (f ∘′ suc) lookup-transpose : ∀ n (ass : List (Stream A)) → @@ -237,33 +237,33 @@ lookup-tails zero as = B.refl lookup-tails (suc n) as = lookup-tails n (as .tail) lookup-evens : ∀ n (as : Stream A) → lookup (evens as) n ≡ lookup as (n * 2) -lookup-evens zero as = P.refl +lookup-evens zero as = ≡.refl lookup-evens (suc n) as = lookup-evens n (as .tail .tail) lookup-odds : ∀ n (as : Stream A) → lookup (odds as) n ≡ lookup as (suc (n * 2)) -lookup-odds zero as = P.refl +lookup-odds zero as = ≡.refl lookup-odds (suc n) as = lookup-odds n (as .tail .tail) lookup-interleave-even : ∀ n (as bs : Stream A) → lookup (interleave as bs) (n * 2) ≡ lookup as n -lookup-interleave-even zero as bs = P.refl +lookup-interleave-even zero as bs = ≡.refl lookup-interleave-even (suc n) as bs = lookup-interleave-even n (as .tail) (bs .tail) lookup-interleave-odd : ∀ n (as bs : Stream A) → lookup (interleave as bs) (suc (n * 2)) ≡ lookup bs n -lookup-interleave-odd zero as bs = P.refl +lookup-interleave-odd zero as bs = ≡.refl lookup-interleave-odd (suc n) as bs = lookup-interleave-odd n (as .tail) (bs .tail) ------------------------------------------------------------------------ -- Properties of take take-iterate : ∀ n f (x : A) → take n (iterate f x) ≡ Vec.iterate f x n -take-iterate zero f x = P.refl +take-iterate zero f x = ≡.refl take-iterate (suc n) f x = cong (x ∷_) (take-iterate n f (f x)) take-zipWith : ∀ n (f : A → B → C) as bs → take n (zipWith f as bs) ≡ Vec.zipWith f (take n as) (take n bs) -take-zipWith zero f as bs = P.refl +take-zipWith zero f as bs = ≡.refl take-zipWith (suc n) f as bs = cong (f (as .head) (bs .head) ∷_) (take-zipWith n f (as .tail) (bs . tail)) @@ -271,21 +271,21 @@ take-zipWith (suc n) f as bs = -- Properties of drop drop-drop : ∀ m n (as : Stream A) → drop n (drop m as) ≡ drop (m + n) as -drop-drop zero n as = P.refl +drop-drop zero n as = ≡.refl drop-drop (suc m) n as = drop-drop m n (as .tail) drop-zipWith : ∀ n (f : A → B → C) as bs → drop n (zipWith f as bs) ≡ zipWith f (drop n as) (drop n bs) -drop-zipWith zero f as bs = P.refl +drop-zipWith zero f as bs = ≡.refl drop-zipWith (suc n) f as bs = drop-zipWith n f (as .tail) (bs .tail) drop-ap : ∀ n (fs : Stream (A → B)) as → drop n (ap fs as) ≡ ap (drop n fs) (drop n as) -drop-ap zero fs as = P.refl +drop-ap zero fs as = ≡.refl drop-ap (suc n) fs as = drop-ap n (fs .tail) (as .tail) drop-iterate : ∀ n f (x : A) → drop n (iterate f x) ≡ iterate f (ℕ.iterate f x n) -drop-iterate zero f x = P.refl +drop-iterate zero f x = ≡.refl drop-iterate (suc n) f x = drop-iterate n f (f x) ------------------------------------------------------------------------ @@ -293,25 +293,25 @@ drop-iterate (suc n) f x = drop-iterate n f (f x) zipWith-defn : ∀ (f : A → B → C) as bs → zipWith f as bs ≈ (repeat f ⟨ ap ⟩ as ⟨ ap ⟩ bs) -zipWith-defn f as bs .head = P.refl +zipWith-defn f as bs .head = ≡.refl zipWith-defn f as bs .tail = zipWith-defn f (as .tail) (bs .tail) zipWith-const : (as : Stream A) (bs : Stream B) → zipWith const as bs ≈ as -zipWith-const as bs .head = P.refl +zipWith-const as bs .head = ≡.refl zipWith-const as bs .tail = zipWith-const (as .tail) (bs .tail) zipWith-flip : ∀ (f : A → B → C) as bs → zipWith (flip f) as bs ≈ zipWith f bs as -zipWith-flip f as bs .head = P.refl +zipWith-flip f as bs .head = ≡.refl zipWith-flip f as bs .tail = zipWith-flip f (as .tail) (bs. tail) ------------------------------------------------------------------------ -- Properties of interleave interleave-evens-odds : (as : Stream A) → interleave (evens as) (odds as) ≈ as -interleave-evens-odds as .head = P.refl -interleave-evens-odds as .tail .head = P.refl +interleave-evens-odds as .head = ≡.refl +interleave-evens-odds as .tail .head = ≡.refl interleave-evens-odds as .tail .tail = interleave-evens-odds (as .tail .tail) ------------------------------------------------------------------------ diff --git a/src/Codata/Musical/Colist/Infinite-merge.agda b/src/Codata/Musical/Colist/Infinite-merge.agda index ecf393ca83..69d07b9523 100644 --- a/src/Codata/Musical/Colist/Infinite-merge.agda +++ b/src/Codata/Musical/Colist/Infinite-merge.agda @@ -25,7 +25,7 @@ open import Function.Related.TypeIsomorphisms open import Level open import Relation.Unary using (Pred) import Induction.WellFounded as WF -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) import Relation.Binary.Construct.On as On @@ -149,9 +149,9 @@ Any-merge {P = P} xss = mk↔ₛ′ (proj₁ ∘ to xss) from to∘from (proj₂ from-injective : ∀ {xss} (p₁ p₂ : Any Q xss) → from p₁ ≡ from p₂ → p₁ ≡ p₂ - from-injective (here (inj₁ p)) (here (inj₁ .p)) P.refl = P.refl + from-injective (here (inj₁ p)) (here (inj₁ .p)) ≡.refl = ≡.refl from-injective (here (inj₂ p₁)) (here (inj₂ p₂)) eq = - P.cong (here ∘ inj₂) $ + ≡.cong (here ∘ inj₂) $ inj₁-injective $ Injection.injective (↔⇒↣ (↔-sym (Any-⋎P _))) $ there-injective eq @@ -166,7 +166,7 @@ Any-merge {P = P} xss = mk↔ₛ′ (proj₁ ∘ to xss) from to∘from (proj₂ (there-injective eq) ... | () from-injective (there {x = _ , xs} p₁) (there p₂) eq = - P.cong there $ + ≡.cong there $ from-injective p₁ p₂ $ inj₂-injective $ Injection.injective (↔⇒↣ (↔-sym (Any-⋎P xs))) $ @@ -188,15 +188,15 @@ Any-merge {P = P} xss = mk↔ₛ′ (proj₁ ∘ to xss) from to∘from (proj₂ step : ∀ p → WF.WfRec (_<′_ on size) InputPred p → InputPred p step ([] , ()) rec - step ((x , xs) ∷ xss , here p) rec = here (inj₁ p) , P.refl + step ((x , xs) ∷ xss , here p) rec = here (inj₁ p) , ≡.refl step ((x , xs) ∷ xss , there p) rec with Inverse.to (Any-⋎P xs) p | Inverse.strictlyInverseʳ (Any-⋎P xs) p | index-Any-⋎P xs p - ... | inj₁ q | P.refl | _ = here (inj₂ q) , P.refl - ... | inj₂ q | P.refl | q≤p = + ... | inj₁ q | ≡.refl | _ = here (inj₂ q) , ≡.refl + ... | inj₂ q | ≡.refl | q≤p = Product.map there - (P.cong (there ∘ (Inverse.from (Any-⋎P xs)) ∘ inj₂)) + (≡.cong (there ∘ (Inverse.from (Any-⋎P xs)) ∘ inj₂)) (rec (s≤′s q≤p)) to∘from = λ p → from-injective _ _ (proj₂ (to xss (from p))) diff --git a/src/Codata/Musical/Conat.agda b/src/Codata/Musical/Conat.agda index f4dab55a6f..760d733a20 100644 --- a/src/Codata/Musical/Conat.agda +++ b/src/Codata/Musical/Conat.agda @@ -14,7 +14,7 @@ open import Function.Base using (_∋_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) ------------------------------------------------------------------------ -- Re-exporting the type and basic operations @@ -27,11 +27,11 @@ open import Codata.Musical.Conat.Base public module Coℕ-injective where suc-injective : ∀ {m n} → (Coℕ ∋ suc m) ≡ suc n → m ≡ n - suc-injective P.refl = P.refl + suc-injective ≡.refl = ≡.refl fromℕ-injective : ∀ {m n} → fromℕ m ≡ fromℕ n → m ≡ n -fromℕ-injective {zero} {zero} eq = P.refl -fromℕ-injective {suc m} {suc n} eq = P.cong suc (fromℕ-injective (P.cong pred eq)) +fromℕ-injective {zero} {zero} eq = ≡.refl +fromℕ-injective {suc m} {suc n} eq = ≡.cong suc (fromℕ-injective (≡.cong pred eq)) ------------------------------------------------------------------------ -- Equality @@ -45,7 +45,7 @@ data _≈_ : Coℕ → Coℕ → Set where module ≈-injective where suc-injective : ∀ {m n p q} → (suc m ≈ suc n ∋ suc p) ≡ suc q → p ≡ q - suc-injective P.refl = P.refl + suc-injective ≡.refl = ≡.refl setoid : Setoid _ _ setoid = record diff --git a/src/Codata/Musical/Covec.agda b/src/Codata/Musical/Covec.agda index 2658fd0311..64156344d9 100644 --- a/src/Codata/Musical/Covec.agda +++ b/src/Codata/Musical/Covec.agda @@ -21,7 +21,7 @@ open import Relation.Binary.Core using (_⇒_; _=[_]⇒_) open import Relation.Binary.Bundles using (Setoid; Poset) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive; Antisymmetric) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) private variable @@ -38,10 +38,10 @@ data Covec (A : Set a) : Coℕ → Set a where _∷_ : ∀ {n} (x : A) (xs : ∞ (Covec A (♭ n))) → Covec A (suc n) ∷-injectiveˡ : ∀ {a b} {n} {as bs} → (Covec A (suc n) ∋ a ∷ as) ≡ b ∷ bs → a ≡ b -∷-injectiveˡ P.refl = P.refl +∷-injectiveˡ ≡.refl = ≡.refl ∷-injectiveʳ : ∀ {a b} {n} {as bs} → (Covec A (suc n) ∋ a ∷ as) ≡ b ∷ bs → as ≡ bs -∷-injectiveʳ P.refl = P.refl +∷-injectiveʳ ≡.refl = ≡.refl ------------------------------------------------------------------------ -- Some operations diff --git a/src/Codata/Musical/Stream.agda b/src/Codata/Musical/Stream.agda index d0d655af65..12bf3b5fb7 100644 --- a/src/Codata/Musical/Stream.agda +++ b/src/Codata/Musical/Stream.agda @@ -16,7 +16,7 @@ open import Data.Nat.Base using (ℕ; zero; suc) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive) open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) private variable @@ -144,13 +144,13 @@ setoid A = record } where refl : Reflexive _≈_ - refl {_ ∷ _} = P.refl ∷ ♯ refl + refl {_ ∷ _} = ≡.refl ∷ ♯ refl sym : Symmetric _≈_ - sym (x≡ ∷ xs≈) = P.sym x≡ ∷ ♯ sym (♭ xs≈) + sym (x≡ ∷ xs≈) = ≡.sym x≡ ∷ ♯ sym (♭ xs≈) trans : Transitive _≈_ - trans (x≡ ∷ xs≈) (y≡ ∷ ys≈) = P.trans x≡ y≡ ∷ ♯ trans (♭ xs≈) (♭ ys≈) + trans (x≡ ∷ xs≈) (y≡ ∷ ys≈) = ≡.trans x≡ y≡ ∷ ♯ trans (♭ xs≈) (♭ ys≈) head-cong : {xs ys : Stream A} → xs ≈ ys → head xs ≡ head ys head-cong (x≡ ∷ _) = x≡ @@ -160,13 +160,13 @@ tail-cong (_ ∷ xs≈) = ♭ xs≈ map-cong : ∀ (f : A → B) {xs ys} → xs ≈ ys → map f xs ≈ map f ys -map-cong f (x≡ ∷ xs≈) = P.cong f x≡ ∷ ♯ map-cong f (♭ xs≈) +map-cong f (x≡ ∷ xs≈) = ≡.cong f x≡ ∷ ♯ map-cong f (♭ xs≈) zipWith-cong : ∀ (_∙_ : A → B → C) {xs xs′ ys ys′} → xs ≈ xs′ → ys ≈ ys′ → zipWith _∙_ xs ys ≈ zipWith _∙_ xs′ ys′ zipWith-cong _∙_ (x≡ ∷ xs≈) (y≡ ∷ ys≈) = - P.cong₂ _∙_ x≡ y≡ ∷ ♯ zipWith-cong _∙_ (♭ xs≈) (♭ ys≈) + ≡.cong₂ _∙_ x≡ y≡ ∷ ♯ zipWith-cong _∙_ (♭ xs≈) (♭ ys≈) infixr 5 _⋎-cong_ diff --git a/src/Codata/Sized/Colist/Bisimilarity.agda b/src/Codata/Sized/Colist/Bisimilarity.agda index 4ea3a18047..94370a7fc2 100644 --- a/src/Codata/Sized/Colist/Bisimilarity.agda +++ b/src/Codata/Sized/Colist/Bisimilarity.agda @@ -20,8 +20,8 @@ open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive; Sym; Trans) -open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_) -import Relation.Binary.PropositionalEquality.Properties as Eq +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +import Relation.Binary.PropositionalEquality.Properties as ≡ private variable @@ -82,16 +82,16 @@ module _ {A : Set a} where _⊢_≈_ = Bisim _≡_ refl : Reflexive (i ⊢_≈_) - refl = reflexive Eq.refl + refl = reflexive ≡.refl fromEq : ∀ {as bs} → as ≡ bs → i ⊢ as ≈ bs - fromEq Eq.refl = refl + fromEq ≡.refl = refl sym : Symmetric (i ⊢_≈_) - sym = symmetric Eq.sym + sym = symmetric ≡.sym trans : Transitive (i ⊢_≈_) - trans = transitive Eq.trans + trans = transitive ≡.trans isEquivalence : {R : Rel A r} → IsEquivalence R → IsEquivalence (Bisim R i) isEquivalence equiv^R = record @@ -107,4 +107,4 @@ setoid S i = record module ≈-Reasoning {a} {A : Set a} {i} where - open import Relation.Binary.Reasoning.Setoid (setoid (Eq.setoid A) i) public + open import Relation.Binary.Reasoning.Setoid (setoid (≡.setoid A) i) public diff --git a/src/Codata/Sized/Covec/Bisimilarity.agda b/src/Codata/Sized/Covec/Bisimilarity.agda index 685b7f180f..f50e249bfa 100644 --- a/src/Codata/Sized/Covec/Bisimilarity.agda +++ b/src/Codata/Sized/Covec/Bisimilarity.agda @@ -15,7 +15,7 @@ open import Codata.Sized.Conat hiding (_⊔_) open import Codata.Sized.Covec open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive; Sym; Trans) -open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) data Bisim {a b r} {A : Set a} {B : Set b} (R : A → B → Set r) (i : Size) : ∀ m n (xs : Covec A ∞ m) (ys : Covec B ∞ n) → Set (r ⊔ a ⊔ b) where @@ -56,10 +56,10 @@ module _ {ℓ} {A : Set ℓ} where _,_⊢_≈_ i m = Bisim _≡_ i m m refl : ∀ {i m} → Reflexive (i , m ⊢_≈_) - refl = reflexive Eq.refl + refl = reflexive ≡.refl sym : ∀ {i m} → Symmetric (i , m ⊢_≈_) - sym = symmetric Eq.sym + sym = symmetric ≡.sym trans : ∀ {i m} → Transitive (i , m ⊢_≈_) - trans = transitive Eq.trans + trans = transitive ≡.trans diff --git a/src/Codata/Sized/Covec/Properties.agda b/src/Codata/Sized/Covec/Properties.agda index 08122efe0b..7d45bad63c 100644 --- a/src/Codata/Sized/Covec/Properties.agda +++ b/src/Codata/Sized/Covec/Properties.agda @@ -14,7 +14,7 @@ open import Codata.Sized.Conat open import Codata.Sized.Covec open import Codata.Sized.Covec.Bisimilarity open import Function.Base using (id; _∘_) -open import Relation.Binary.PropositionalEquality.Core as Eq +open import Relation.Binary.PropositionalEquality.Core as ≡ -- Functor laws @@ -22,13 +22,13 @@ module _ {a} {A : Set a} where map-id : ∀ {m} (as : Covec A ∞ m) {i} → i , m ⊢ 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) module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where map-∘ : ∀ (f : A → B) (g : B → C) {m} as {i} → i , m ⊢ 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) + map-∘ f g (a ∷ as) = ≡.refl ∷ λ where .force → map-∘ f g (as .force) ------------------------------------------------------------------------ -- DEPRECATED diff --git a/src/Codata/Sized/Cowriter/Bisimilarity.agda b/src/Codata/Sized/Cowriter/Bisimilarity.agda index d9ce3e2d6f..c38ef8180b 100644 --- a/src/Codata/Sized/Cowriter/Bisimilarity.agda +++ b/src/Codata/Sized/Cowriter/Bisimilarity.agda @@ -17,8 +17,8 @@ open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive; Sym; Trans) -open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_) -import Relation.Binary.PropositionalEquality.Properties as Eq +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +import Relation.Binary.PropositionalEquality.Properties as ≡ private variable @@ -74,16 +74,16 @@ module _ {W : Set w} {A : Set a} where _⊢_≈_ = Bisim _≡_ _≡_ refl : Reflexive (i ⊢_≈_) - refl = reflexive Eq.refl Eq.refl + refl = reflexive ≡.refl ≡.refl fromEq : ∀ {as bs} → as ≡ bs → i ⊢ as ≈ bs - fromEq Eq.refl = refl + fromEq ≡.refl = refl sym : Symmetric (i ⊢_≈_) - sym = symmetric Eq.sym Eq.sym + sym = symmetric ≡.sym ≡.sym trans : Transitive (i ⊢_≈_) - trans = transitive Eq.trans Eq.trans + trans = transitive ≡.trans ≡.trans module _ {R : Rel W r} {S : Rel A s} (equiv^R : IsEquivalence R) (equiv^S : IsEquivalence S) where @@ -108,4 +108,4 @@ setoid R S i = record module ≈-Reasoning {W : Set w} {A : Set a} {i} where open import Relation.Binary.Reasoning.Setoid - (setoid (Eq.setoid W) (Eq.setoid A) i) public + (setoid (≡.setoid W) (≡.setoid A) i) public diff --git a/src/Codata/Sized/Delay/Bisimilarity.agda b/src/Codata/Sized/Delay/Bisimilarity.agda index f4845669ba..fb01d94ebc 100644 --- a/src/Codata/Sized/Delay/Bisimilarity.agda +++ b/src/Codata/Sized/Delay/Bisimilarity.agda @@ -14,7 +14,7 @@ open import Codata.Sized.Delay open import Level open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive; Sym; Trans) -open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) data Bisim {a b r} {A : Set a} {B : Set b} (R : A → B → Set r) i : (xs : Delay A ∞) (ys : Delay B ∞) → Set (a ⊔ b ⊔ r) where @@ -53,10 +53,10 @@ module _ {ℓ} {A : Set ℓ} where _⊢_≈_ = Bisim _≡_ refl : ∀ {i} → Reflexive (i ⊢_≈_) - refl = reflexive Eq.refl + refl = reflexive ≡.refl sym : ∀ {i} → Symmetric (i ⊢_≈_) - sym = symmetric Eq.sym + sym = symmetric ≡.sym trans : ∀ {i} → Transitive (i ⊢_≈_) - trans = transitive Eq.trans + trans = transitive ≡.trans diff --git a/src/Codata/Sized/Delay/Properties.agda b/src/Codata/Sized/Delay/Properties.agda index 39a56d1ef9..de2b3c95b4 100644 --- a/src/Codata/Sized/Delay/Properties.agda +++ b/src/Codata/Sized/Delay/Properties.agda @@ -17,7 +17,7 @@ open import Codata.Sized.Conat.Bisimilarity as Coℕ using (zero ; suc) open import Codata.Sized.Delay open import Codata.Sized.Delay.Bisimilarity open import Function.Base using (id; _∘′_) -open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) module _ {a} {A : Set a} where @@ -40,19 +40,19 @@ module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where suc λ where .force → length-zipWith f (da .force) (db .force) map-id : ∀ da {i} → i ⊢ map (id {A = A}) da ≈ da - map-id (now a) = now Eq.refl + map-id (now a) = now ≡.refl map-id (later da) = later λ where .force → map-id (da .force) map-∘ : ∀ (f : A → B) (g : B → C) da {i} → i ⊢ map g (map f da) ≈ map (g ∘′ f) da - map-∘ f g (now a) = now Eq.refl + map-∘ f g (now a) = now ≡.refl map-∘ f g (later da) = later λ where .force → map-∘ f g (da .force) map-unfold : ∀ (f : B → C) n (s : A) {i} → i ⊢ map f (unfold n s) ≈ unfold (Sum.map id f ∘′ n) s map-unfold f n s with n s ... | Sum.inj₁ s′ = later λ where .force → map-unfold f n s′ - ... | Sum.inj₂ b = now Eq.refl + ... | Sum.inj₂ b = now ≡.refl ------------------------------------------------------------------------ @@ -62,9 +62,9 @@ module _ {a b c} {A : Set a} {B : Set b} {C : Set c} where {d : Delay A ∞} → (d⇓₁ : d ⇓) → (d⇓₂ : d ⇓) → d⇓₁ ≡ d⇓₂ -⇓-unique {d = now s} (now s) (now s) = Eq.refl +⇓-unique {d = now s} (now s) (now s) = ≡.refl ⇓-unique {d = later d'} (later l) (later r) = - Eq.cong later (⇓-unique {d = force d'} l r) + ≡.cong later (⇓-unique {d = force d'} l r) module _ {a} {A B : Set a} where @@ -85,7 +85,7 @@ module _ {a} {A B : Set a} where extract-bind-⇓ : {d : Delay A Size.∞} → {f : A → Delay B Size.∞} → (d⇓ : d ⇓) → (f⇓ : f (extract d⇓) ⇓) → extract (bind-⇓ d⇓ {f} f⇓) ≡ extract f⇓ - extract-bind-⇓ (now a) f⇓ = Eq.refl + extract-bind-⇓ (now a) f⇓ = ≡.refl extract-bind-⇓ (later t) f⇓ = extract-bind-⇓ t f⇓ -- If the right element of a bind returns a certain value so does the @@ -94,7 +94,7 @@ module _ {a} {A B : Set a} where (d : Delay A ∞) {f : A → Delay B ∞} → (bind⇓ : bind d f ⇓) → extract (bind̅₂ d bind⇓) ≡ extract bind⇓ - extract-bind̅₂-bind⇓ (now s) bind⇓ = Eq.refl + extract-bind̅₂-bind⇓ (now s) bind⇓ = ≡.refl extract-bind̅₂-bind⇓ (later s) (later bind⇓) = extract-bind̅₂-bind⇓ (force s) bind⇓ @@ -106,9 +106,9 @@ module _ {a} {A B : Set a} where (d⇓ : d ⇓) → (f⇓ : f (extract d⇓) ⇓) → toℕ (length-⇓ bind⇓) ≡ toℕ (length-⇓ d⇓) ℕ.+ toℕ (length-⇓ f⇓) bind⇓-length {f = f} bind⇓ d⇓@(now s') f⇓ = - Eq.cong (toℕ ∘′ length-⇓) (⇓-unique bind⇓ f⇓) + ≡.cong (toℕ ∘′ length-⇓) (⇓-unique bind⇓ f⇓) bind⇓-length {d = d@(later dt)} {f = f} bind⇓@(later bind'⇓) d⇓@(later r) f⇓ = - Eq.cong ℕ.suc (bind⇓-length bind'⇓ r f⇓) + ≡.cong ℕ.suc (bind⇓-length bind'⇓ r f⇓) ------------------------------------------------------------------------ -- DEPRECATED diff --git a/src/Codata/Sized/M/Bisimilarity.agda b/src/Codata/Sized/M/Bisimilarity.agda index 9c641ec834..cc0b934990 100644 --- a/src/Codata/Sized/M/Bisimilarity.agda +++ b/src/Codata/Sized/M/Bisimilarity.agda @@ -21,7 +21,7 @@ open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive) -import Relation.Binary.PropositionalEquality.Core as P +import Relation.Binary.PropositionalEquality.Core as ≡ data Bisim {s p} (C : Container s p) (i : Size) : Rel (M C ∞) (s ⊔ p) where inf : ∀ {t u} → Pointwise C (Thunk^R (Bisim C) i) t u → Bisim C i (inf t) (inf u) @@ -32,14 +32,14 @@ module _ {s p} {C : Container s p} where -- combinators C.refl, C.sym and C.trans refl : ∀ {i} → Reflexive (Bisim C i) - refl {x = inf t} = inf (P.refl , λ where p .force → refl) + refl {x = inf t} = inf (≡.refl , λ where p .force → refl) sym : ∀ {i} → Symmetric (Bisim C i) - sym (inf (P.refl , f)) = inf (P.refl , λ where p .force → sym (f p .force)) + sym (inf (≡.refl , f)) = inf (≡.refl , λ where p .force → sym (f p .force)) trans : ∀ {i} → Transitive (Bisim C i) - trans (inf (P.refl , f)) (inf (P.refl , g)) = - inf (P.refl , λ where p .force → trans (f p .force) (g p .force)) + trans (inf (≡.refl , f)) (inf (≡.refl , g)) = + inf (≡.refl , λ where p .force → trans (f p .force) (g p .force)) isEquivalence : ∀ {i} → IsEquivalence (Bisim C i) isEquivalence = record diff --git a/src/Codata/Sized/M/Properties.agda b/src/Codata/Sized/M/Properties.agda index a02410d43e..8d9724f9a7 100644 --- a/src/Codata/Sized/M/Properties.agda +++ b/src/Codata/Sized/M/Properties.agda @@ -18,19 +18,19 @@ import Data.Container.Morphism as Mp 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 -import Relation.Binary.PropositionalEquality.Properties as P +import Relation.Binary.PropositionalEquality.Core as ≡ +import Relation.Binary.PropositionalEquality.Properties as ≡ open import Data.Container.Relation.Binary.Pointwise using (_,_) import Data.Container.Relation.Binary.Equality.Setoid as EqSetoid -private module Eq {a} (A : Set a) = EqSetoid (P.setoid A) +private module Eq {a} (A : Set a) = EqSetoid (≡.setoid A) open Eq using (Eq) module _ {s p} {C : Container s p} where map-id : ∀ {i} c → Bisim C i (map (Mp.id C) c) c - map-id (inf (s , f)) = inf (P.refl , λ where p .force → map-id (f p .force)) + map-id (inf (s , f)) = inf (≡.refl , λ where p .force → map-id (f p .force)) module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂} where @@ -39,7 +39,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s ∀ c₁ → Bisim C₂ i (map f c₁) (map g c₁) map-cong {f = f} {g} f≗g (inf t@(s , n)) with f≗g t ... | eqs , eqf = inf (eqs , λ where - p .force {j} → P.subst (λ t → Bisim C₂ j (map f (n (position f p) .force)) + p .force {j} → ≡.subst (λ t → Bisim C₂ j (map f (n (position f p) .force)) (map g (t .force))) (eqf p) (map-cong f≗g (n (position f p) .force))) @@ -49,7 +49,7 @@ module _ {s₁ s₂ s₃ p₁ p₂ p₃} {C₁ : Container s₁ p₁} map-∘ : ∀ {i} {g : C₂ ⇒ C₃} {f : C₁ ⇒ C₂} c₁ → Bisim C₃ i (map (g Mp.∘ f) c₁) (map g $′ map f c₁) - map-∘ (inf (s , f)) = inf (P.refl , λ where p .force → map-∘ (f _ .force)) + map-∘ (inf (s , f)) = inf (≡.refl , λ where p .force → map-∘ (f _ .force)) module _ {s₁ s₂ p₁ p₂ s} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂} @@ -57,7 +57,7 @@ module _ {s₁ s₂ p₁ p₂ s} {C₁ : Container s₁ p₁} {C₂ : Container map-unfold : ∀ {i} s → Bisim C₂ i (map f (unfold alg s)) (unfold (⟪ f ⟫ ∘′ alg) s) - map-unfold s = inf (P.refl , λ where p .force → map-unfold _) + map-unfold s = inf (≡.refl , λ where p .force → map-unfold _) ------------------------------------------------------------------------ -- DEPRECATED diff --git a/src/Codata/Sized/Stream/Bisimilarity.agda b/src/Codata/Sized/Stream/Bisimilarity.agda index f08d0daeb8..57516709df 100644 --- a/src/Codata/Sized/Stream/Bisimilarity.agda +++ b/src/Codata/Sized/Stream/Bisimilarity.agda @@ -19,8 +19,8 @@ open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive; Sym; Trans) open import Relation.Binary.Structures using (IsEquivalence) -open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_) -import Relation.Binary.PropositionalEquality.Properties as Eq +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +import Relation.Binary.PropositionalEquality.Properties as ≡ private variable @@ -87,14 +87,14 @@ module _ {A : Set a} where _⊢_≈_ = Bisim _≡_ refl : ∀ {i} → Reflexive (i ⊢_≈_) - refl = reflexive Eq.refl + refl = reflexive ≡.refl sym : ∀ {i} → Symmetric (i ⊢_≈_) - sym = symmetric Eq.sym + sym = symmetric ≡.sym trans : ∀ {i} → Transitive (i ⊢_≈_) - trans = transitive Eq.trans + trans = transitive ≡.trans module ≈-Reasoning {a} {A : Set a} {i} where - open import Relation.Binary.Reasoning.Setoid (setoid (Eq.setoid A) i) public + open import Relation.Binary.Reasoning.Setoid (setoid (≡.setoid A) i) public diff --git a/src/Codata/Sized/Stream/Properties.agda b/src/Codata/Sized/Stream/Properties.agda index 0fd436f1bd..913b84edae 100644 --- a/src/Codata/Sized/Stream/Properties.agda +++ b/src/Codata/Sized/Stream/Properties.agda @@ -19,12 +19,12 @@ 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 +import Data.List.Relation.Binary.Equality.Propositional as ≋ open import Data.Product.Base as Product using (_,_) open import Data.Vec.Base as Vec using (_∷_) open import Function.Base using (id; _$_; _∘′_; const) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; _≢_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; _≢_) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) @@ -40,63 +40,63 @@ private -- repeat lookup-repeat-identity : (n : ℕ) (a : A) → lookup (repeat a) n ≡ a -lookup-repeat-identity zero a = P.refl +lookup-repeat-identity zero a = ≡.refl lookup-repeat-identity (suc n) a = lookup-repeat-identity n a take-repeat-identity : (n : ℕ) (a : A) → take n (repeat a) ≡ Vec.replicate n a -take-repeat-identity zero a = P.refl -take-repeat-identity (suc n) a = P.cong (a Vec.∷_) (take-repeat-identity n a) +take-repeat-identity zero a = ≡.refl +take-repeat-identity (suc n) a = ≡.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 (Product.map₁ (a ∷_)) (splitAt-repeat-identity n a) +splitAt-repeat-identity zero a = ≡.refl +splitAt-repeat-identity (suc n) a = ≡.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 -replicate-repeat (suc n) a = P.refl ∷ λ where .force → replicate-repeat n a +replicate-repeat (suc n) a = ≡.refl ∷ λ where .force → replicate-repeat n a cycle-replicate : ∀ {i} (n : ℕ) (n≢0 : n ≢ 0) (a : A) → i ⊢ cycle (List⁺.replicate n n≢0 a) ≈ repeat a cycle-replicate {i} n n≢0 a = let as = List⁺.replicate n n≢0 a in begin cycle as ≡⟨⟩ - as ⁺++ _ ≈⟨ ⁺++⁺ Eq.≋-refl (λ where .force → cycle-replicate n n≢0 a) ⟩ - as ⁺++ (λ where .force → repeat a) ≈⟨ P.refl ∷ (λ where .force → replicate-repeat (pred n) a) ⟩ + as ⁺++ _ ≈⟨ ⁺++⁺ ≋.≋-refl (λ where .force → cycle-replicate n n≢0 a) ⟩ + as ⁺++ (λ where .force → repeat a) ≈⟨ ≡.refl ∷ (λ where .force → replicate-repeat (pred n) a) ⟩ repeat a ∎ where open ≈-Reasoning module _ {a b} {A : Set a} {B : Set b} where map-repeat : ∀ (f : A → B) a {i} → i ⊢ map f (repeat a) ≈ repeat (f a) - map-repeat f a = P.refl ∷ λ where .force → map-repeat f a + map-repeat f a = ≡.refl ∷ λ where .force → map-repeat f a ap-repeat : ∀ (f : A → B) a {i} → i ⊢ ap (repeat f) (repeat a) ≈ repeat (f a) - ap-repeat f a = P.refl ∷ λ where .force → ap-repeat f a + ap-repeat f a = ≡.refl ∷ λ where .force → ap-repeat f a ap-repeatˡ : ∀ (f : A → B) as {i} → i ⊢ ap (repeat f) as ≈ map f as - ap-repeatˡ f (a ∷ as) = P.refl ∷ λ where .force → ap-repeatˡ f (as .force) + ap-repeatˡ f (a ∷ as) = ≡.refl ∷ λ where .force → ap-repeatˡ f (as .force) ap-repeatʳ : ∀ (fs : Stream (A → B) ∞) (a : A) {i} → i ⊢ ap fs (repeat a) ≈ map (_$ a) fs - ap-repeatʳ (f ∷ fs) a = P.refl ∷ λ where .force → ap-repeatʳ (fs .force) a + ap-repeatʳ (f ∷ fs) a = ≡.refl ∷ λ where .force → ap-repeatʳ (fs .force) a map-++ : ∀ {i} (f : A → B) as xs → i ⊢ map f (as ++ xs) ≈ List.map f as ++ map f xs map-++ f [] xs = refl - map-++ f (a ∷ as) xs = P.refl ∷ λ where .force → map-++ f as xs + map-++ f (a ∷ as) xs = ≡.refl ∷ λ where .force → map-++ f as xs map-⁺++ : ∀ {i} (f : A → B) as xs → i ⊢ map f (as ⁺++ xs) ≈ List⁺.map f as ⁺++ Thunk.map (map f) xs - map-⁺++ f (a ∷ as) xs = P.refl ∷ (λ where .force → map-++ f as (xs .force)) + map-⁺++ f (a ∷ as) xs = ≡.refl ∷ (λ where .force → map-++ f as (xs .force)) map-cycle : ∀ {i} (f : A → B) as → i ⊢ map f (cycle as) ≈ cycle (List⁺.map f as) map-cycle f as = begin map f (cycle as) ≈⟨ map-⁺++ f as _ ⟩ - List⁺.map f as ⁺++ _ ≈⟨ ⁺++⁺ Eq.≋-refl (λ where .force → map-cycle f as) ⟩ + List⁺.map f as ⁺++ _ ≈⟨ ⁺++⁺ ≋.≋-refl (λ where .force → map-cycle f as) ⟩ cycle (List⁺.map f as) ∎ where open ≈-Reasoning ------------------------------------------------------------------------ -- Functor laws map-id : ∀ (as : Stream A ∞) → i ⊢ map id as ≈ as -map-id (a ∷ as) = P.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 ⊢ map g (map f as) ≈ map (g ∘′ f) as -map-∘ f g (a ∷ as) = P.refl ∷ λ where .force → map-∘ f g (as .force) +map-∘ f g (a ∷ as) = ≡.refl ∷ λ where .force → map-∘ f g (as .force) ------------------------------------------------------------------------ @@ -104,19 +104,19 @@ map-∘ f g (a ∷ as) = P.refl ∷ λ where .force → map-∘ f g (as .force) splitAt-map : ∀ n (f : A → B) 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 zero f xs = ≡.refl splitAt-map (suc n) f (x ∷ xs) = - P.cong (Product.map₁ (f x Vec.∷_)) (splitAt-map n f (xs .force)) + ≡.cong (Product.map₁ (f x Vec.∷_)) (splitAt-map n f (xs .force)) ------------------------------------------------------------------------ -- iterate lookup-iterate-identity : ∀ n f (a : A) → lookup (iterate f a) n ≡ fold a f n -lookup-iterate-identity zero f a = P.refl +lookup-iterate-identity zero f a = ≡.refl lookup-iterate-identity (suc n) f a = begin lookup (iterate f a) (suc n) ≡⟨⟩ lookup (iterate f (f a)) n ≡⟨ lookup-iterate-identity n f (f a) ⟩ - fold (f a) f n ≡⟨ fold-pull a f (const ∘′ f) (f a) P.refl (λ _ → P.refl) n ⟩ + fold (f a) f n ≡⟨ fold-pull a f (const ∘′ f) (f a) ≡.refl (λ _ → ≡.refl) n ⟩ f (fold a f n) ≡⟨⟩ fold a f (suc n) ∎ where open ≡-Reasoning diff --git a/src/Data/List/Relation/Binary/Equality/Propositional.agda b/src/Data/List/Relation/Binary/Equality/Propositional.agda index 8b45be7d27..2b458c29b9 100644 --- a/src/Data/List/Relation/Binary/Equality/Propositional.agda +++ b/src/Data/List/Relation/Binary/Equality/Propositional.agda @@ -16,20 +16,20 @@ module Data.List.Relation.Binary.Equality.Propositional {a} {A : Set a} where open import Data.List.Base import Data.List.Relation.Binary.Equality.Setoid as SetoidEquality -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 ≡ ------------------------------------------------------------------------ -- Re-export everything from setoid equality -open SetoidEquality (P.setoid A) public +open SetoidEquality (≡.setoid A) public ------------------------------------------------------------------------ -- ≋ is propositional ≋⇒≡ : _≋_ ⇒ _≡_ -≋⇒≡ [] = P.refl -≋⇒≡ (P.refl ∷ xs≈ys) = P.cong (_ ∷_) (≋⇒≡ xs≈ys) +≋⇒≡ [] = ≡.refl +≋⇒≡ (≡.refl ∷ xs≈ys) = ≡.cong (_ ∷_) (≋⇒≡ xs≈ys) ≡⇒≋ : _≡_ ⇒ _≋_ -≡⇒≋ P.refl = ≋-refl +≡⇒≋ ≡.refl = ≋-refl diff --git a/src/Effect/Applicative.agda b/src/Effect/Applicative.agda index 95e2debda8..c681a18133 100644 --- a/src/Effect/Applicative.agda +++ b/src/Effect/Applicative.agda @@ -21,7 +21,7 @@ open import Effect.Functor as Fun using (RawFunctor) open import Function.Base using (const; flip; _∘′_) open import Level using (Level; suc; _⊔_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) private variable diff --git a/src/Effect/Applicative/Indexed.agda b/src/Effect/Applicative/Indexed.agda index f547e68b27..1aa639292a 100644 --- a/src/Effect/Applicative/Indexed.agda +++ b/src/Effect/Applicative/Indexed.agda @@ -15,7 +15,7 @@ open import Effect.Functor using (RawFunctor) open import Data.Product.Base using (_×_; _,_) open import Function.Base open import Level -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) @@ -111,6 +111,6 @@ record Morphism {I : Set i} {F₁ F₂ : IFun I f} op (f A₁.<$> x) ≡ (f A₂.<$> op x) op-<$> f x = begin op (A₁._⊛_ (A₁.pure f) x) ≡⟨ op-⊛ _ _ ⟩ - A₂._⊛_ (op (A₁.pure f)) (op x) ≡⟨ P.cong₂ A₂._⊛_ (op-pure _) P.refl ⟩ + A₂._⊛_ (op (A₁.pure f)) (op x) ≡⟨ ≡.cong₂ A₂._⊛_ (op-pure _) ≡.refl ⟩ A₂._⊛_ (A₂.pure f) (op x) ∎ where open ≡-Reasoning diff --git a/src/Effect/Monad/Partiality.agda b/src/Effect/Monad/Partiality.agda index 44cf13464d..c0723f1370 100644 --- a/src/Effect/Monad/Partiality.agda +++ b/src/Effect/Monad/Partiality.agda @@ -27,8 +27,8 @@ open import Relation.Binary.Structures open import Relation.Binary.Bundles using (Preorder; Setoid; Poset) import Relation.Binary.Properties.Setoid as SetoidProperties -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 ≡ open import Relation.Nullary open import Relation.Nullary.Decidable hiding (map) open import Relation.Nullary.Negation @@ -117,13 +117,13 @@ data Kind : Set where infix 4 _≟-Kind_ _≟-Kind_ : Decidable (_≡_ {A = Kind}) -_≟-Kind_ strong strong = yes P.refl +_≟-Kind_ strong strong = yes ≡.refl _≟-Kind_ strong (other k) = no λ() _≟-Kind_ (other k) strong = no λ() -_≟-Kind_ (other geq) (other geq) = yes P.refl +_≟-Kind_ (other geq) (other geq) = yes ≡.refl _≟-Kind_ (other geq) (other weak) = no λ() _≟-Kind_ (other weak) (other geq) = no λ() -_≟-Kind_ (other weak) (other weak) = yes P.refl +_≟-Kind_ (other weak) (other weak) = yes ≡.refl -- A predicate which is satisfied only for equalities. Note that, for -- concrete inputs, this predicate evaluates to ⊤ or ⊥. @@ -303,14 +303,14 @@ module _ {A : Set a} {_∼_ : A → A → Set ℓ} where ; _≈_ = _≡_ ; _≲_ = Rel k ; isPreorder = record - { isEquivalence = P.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = refl′ ; trans = Equivalence.trans (IsPreorder.trans pre) } } where refl′ : ∀ {k} {x y : A ⊥} → x ≡ y → Rel k x y - refl′ P.refl = Equivalence.refl (IsPreorder.refl pre) + refl′ ≡.refl = Equivalence.refl (IsPreorder.refl pre) private preorder′ : IsEquivalence _∼_ → Kind → Preorder _ _ _ @@ -371,7 +371,7 @@ module _ {A : Set a} {_∼_ : A → A → Set ℓ} where infixr 2 _≡⟨_⟩_ _≅⟨_⟩_ _≳⟨_⟩_ _≈⟨_⟩_ _≡⟨_⟩_ : ∀ {k} x {y z : A ⊥} → x ≡ y → Rel k y z → Rel k x z - _ ≡⟨ P.refl ⟩ y∼z = y∼z + _ ≡⟨ ≡.refl ⟩ y∼z = y∼z _≅⟨_⟩_ : ∀ {k} x {y z : A ⊥} → x ≅ y → Rel k y z → Rel k x z _ ≅⟨ x≅y ⟩ y∼z = Pre.trans (≅⇒ x≅y) y∼z @@ -436,7 +436,7 @@ module _ {A : Set a} {_∼_ : A → A → Set ℓ} where ≡⇒ : Reflexive _∼_ → ∀ {k x y} → Equality.Rel _≡_ k x y → Rel k x y - ≡⇒ refl-∼ = map (flip (P.subst (_∼_ _)) refl-∼) + ≡⇒ refl-∼ = map (flip (≡.subst (_∼_ _)) refl-∼) ------------------------------------------------------------------------ -- Steps @@ -454,17 +454,17 @@ module _ {A : Set a} {_∼_ : A → A → Set ℓ} where ∀ {k x y} {z : A} (x≅y : x ≅ y) (y⇓z : y ⇓[ k ] z) → steps (Equivalence.trans trans-∼ (≅⇒ x≅y) y⇓z) ≡ steps y⇓z - left-identity (now _) (now _) = P.refl + left-identity (now _) (now _) = ≡.refl left-identity (later x≅y) (laterˡ y⇓z) = - P.cong suc $ left-identity (♭ x≅y) y⇓z + ≡.cong suc $ left-identity (♭ x≅y) y⇓z right-identity : ∀ {k x} {y z : A} (x⇓y : x ⇓[ k ] y) (y≈z : now y ⇓[ k ] z) → steps (Equivalence.trans trans-∼ x⇓y y≈z) ≡ steps x⇓y - right-identity (now x∼y) (now y∼z) = P.refl + right-identity (now x∼y) (now y∼z) = ≡.refl right-identity (laterˡ x∼y) (now y∼z) = - P.cong suc $ right-identity x∼y (now y∼z) + ≡.cong suc $ right-identity x∼y (now y∼z) ------------------------------------------------------------------------ -- Laws related to bind @@ -538,9 +538,9 @@ module _ {A B : Set s} ∃ λ z → ∃₂ λ (x⇓ : x ⇓[ k ]A z) (fz⇓ : f z ⇓[ k ]B y) → steps x⇓ + steps fz⇓ ≡ steps x>>=f⇓ >>=-inversion-⇓ refl (now x) fx⇓ = - (x , now refl , fx⇓ , P.refl) + (x , now refl , fx⇓ , ≡.refl) >>=-inversion-⇓ refl (later x) (laterˡ x>>=f⇓) = - Prod.map id (Prod.map laterˡ (Prod.map id (P.cong suc))) $ + Prod.map id (Prod.map laterˡ (Prod.map id (≡.cong suc))) $ >>=-inversion-⇓ refl (♭ x) x>>=f⇓ >>=-inversion-⇑ : IsEquivalence _∼A_ → @@ -589,7 +589,7 @@ module _ {A B : Set ℓ} {_∼_ : B → B → Set ℓ} where Rel _∼_ k (x₁ >>= f₁) (x₂ >>= f₂) _≡->>=-cong_ {k} {f₁ = f₁} {f₂} x₁≈x₂ f₁≈f₂ = x₁≈x₂ >>=-cong λ {x} x≡x′ → - P.subst (λ y → Rel _∼_ k (f₁ x) (f₂ y)) x≡x′ (f₁≈f₂ x) + ≡.subst (λ y → Rel _∼_ k (f₁ x) (f₂ y)) x≡x′ (f₁≈f₂ x) ------------------------------------------------------------------------ -- Productivity checker workaround @@ -641,7 +641,7 @@ module Workaround {a} where private open module Eq {A : Set a} = Equality {A = A} _≡_ - open module R {A : Set a} = Reasoning (P.isEquivalence {A = A}) + open module R {A : Set a} = Reasoning (≡.isEquivalence {A = A}) now-hom : (x : A) → ⟦ now x ⟧P ≅ now x now-hom x = now x ∎ @@ -820,7 +820,7 @@ module AlternativeEquality {a ℓ} where whnf≅ (x₁≅x₂ >>= f₁≅f₂) = whnf≅ x₁≅x₂ >>=W λ xRy → whnf≅ (f₁≅f₂ xRy) whnf≅ (x ∎) = reflW x whnf≅ (sym x≅y) = symW _ (whnf≅ x≅y) - whnf≅ (x ≡⟨ P.refl ⟩ y≅z) = whnf≅ y≅z + whnf≅ (x ≡⟨ ≡.refl ⟩ y≅z) = whnf≅ y≅z whnf≅ (x ≅⟨ x≅y ⟩ y≅z) = trans≅W (whnf≅ x≅y) (whnf≅ y≅z) -- More transitivity lemmas. @@ -853,7 +853,7 @@ module AlternativeEquality {a ℓ} where whnf≳ (laterˡ x≲y) = laterˡ (whnf≳ x≲y) whnf≳ (x₁∼x₂ >>= f₁∼f₂) = whnf≳ x₁∼x₂ >>=W λ xRy → whnf≳ (f₁∼f₂ xRy) whnf≳ (x ∎) = reflW x - whnf≳ (x ≡⟨ P.refl ⟩ y≳z) = whnf≳ y≳z + whnf≳ (x ≡⟨ ≡.refl ⟩ y≳z) = whnf≳ y≳z whnf≳ (x ≅⟨ x≅y ⟩ y≳z) = trans≅∼W (whnf≅ x≅y) (whnf≳ y≳z) whnf≳ (x ≳⟨ x≳y ⟩ y≳z) = trans≳-W x≳y (whnf≳ y≳z) whnf≳ (x ≳⟨ x≳y ⟩≅ y≅z) = trans∼≅W (whnf≳ x≳y) (whnf≅ y≅z) @@ -879,7 +879,7 @@ module AlternativeEquality {a ℓ} where whnf (x₁∼x₂ >>= f₁∼f₂) = whnf x₁∼x₂ >>=W λ xRy → whnf (f₁∼f₂ xRy) whnf (x ∎) = reflW x whnf (sym {eq = eq} x≈y) = symW eq (whnf x≈y) - whnf (x ≡⟨ P.refl ⟩ y∼z) = whnf y∼z + whnf (x ≡⟨ ≡.refl ⟩ y∼z) = whnf y∼z whnf (x ≅⟨ x≅y ⟩ y∼z) = trans≅∼W (whnf x≅y) (whnf y∼z) whnf (x ≳⟨ x≳y ⟩ y≳z) = trans≳-W x≳y (whnf y≳z) whnf (x ≳⟨ x≳y ⟩≅ y≅z) = trans∼≅W (whnf x≳y) (whnf y≅z) @@ -930,7 +930,7 @@ idempotent {A = A} B x f = sound (idem x) (x >>= λ y′ → f y′ y′) idem (now x) = f x x ∎ idem (later x) = later (♯ ( - (♭ x >>= λ y′ → later x >>= λ y″ → f y′ y″) ≳⟨ (refl P.refl {x = ♭ x} ≡->>=-cong λ _ → + (♭ x >>= λ y′ → later x >>= λ y″ → f y′ y″) ≳⟨ (refl ≡.refl {x = ♭ x} ≡->>=-cong λ _ → laterˡ (refl (Setoid.refl B))) ⟩ (♭ x >>= λ y′ → ♭ x >>= λ y″ → f y′ y″) ≳⟨ idem (♭ x) ⟩≅ (♭ x >>= λ y′ → f y′ y′) ∎)) diff --git a/src/Effect/Monad/Partiality/All.agda b/src/Effect/Monad/Partiality/All.agda index c5bf2da9ef..6f7b30ef41 100644 --- a/src/Effect/Monad/Partiality/All.agda +++ b/src/Effect/Monad/Partiality/All.agda @@ -15,7 +15,7 @@ open import Function.Base using (flip; _∘_) open import Level open import Relation.Binary.Definitions using (_Respects_) open import Relation.Binary.Structures using (IsEquivalence) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open Partiality._⊥ open Partiality.Equality using (Rel) @@ -82,7 +82,7 @@ module Reasoning {P : A → Set p} infixr 2 _≡⟨_⟩_ _∼⟨_⟩_ _≡⟨_⟩_ : ∀ x {y} → x ≡ y → All P y → All P x - _ ≡⟨ P.refl ⟩ p = p + _ ≡⟨ ≡.refl ⟩ p = p _∼⟨_⟩_ : ∀ {k} x {y} → Rel _∼_ k x y → All P y → All P x _ ∼⟨ x∼y ⟩ p = respects-flip resp (⇒≈ x∼y) p @@ -98,7 +98,7 @@ module Reasoning {P : A → Set p} -- equality. module Reasoning-≡ {a p} {A : Set a} {P : A → Set p} - = Reasoning {P = P} {_∼_ = _≡_} (P.subst P ∘ P.sym) + = Reasoning {P = P} {_∼_ = _≡_} (≡.subst P ∘ ≡.sym) ------------------------------------------------------------------------ -- An alternative, but equivalent, formulation of All @@ -140,12 +140,12 @@ module Alternative {a p : Level} where trans-≅ : {P : A → Set p} {x y : A ⊥} → x ≅ y → AllW P y → AllW P x - trans-≅ (now P.refl) (now p) = now p + trans-≅ (now ≡.refl) (now p) = now p trans-≅ (later x≅y) (later p) = later (_ ≅⟨ ♭ x≅y ⟩P p) trans-≳ : {P : A → Set p} {x y : A ⊥} → x ≳ y → AllW P y → AllW P x - trans-≳ (now P.refl) (now p) = now p + trans-≳ (now ≡.refl) (now p) = now p trans-≳ (later x≳y) (later p) = later (_ ≳⟨ ♭ x≳y ⟩P p) trans-≳ (laterˡ x≳y) p = later (_ ≳⟨ x≳y ⟩P program p) diff --git a/src/Function/Endomorphism/Propositional.agda b/src/Function/Endomorphism/Propositional.agda index c41305f9e9..1846061b17 100644 --- a/src/Function/Endomorphism/Propositional.agda +++ b/src/Function/Endomorphism/Propositional.agda @@ -21,10 +21,10 @@ open import Data.Product.Base using (_,_) open import Function.Base using (id; _∘′_; _∋_) open import Function.Equality using (_⟨$⟩_) open import Relation.Binary.Core using (_Preserves_⟶_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; refl) -import Relation.Binary.PropositionalEquality.Properties as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +import Relation.Binary.PropositionalEquality.Properties as ≡ -import Function.Endomorphism.Setoid (P.setoid A) as Setoid +import Function.Endomorphism.Setoid (≡.setoid A) as Setoid Endo : Set a Endo = A → A @@ -38,7 +38,7 @@ fromSetoidEndo = _⟨$⟩_ toSetoidEndo : Endo → Setoid.Endo toSetoidEndo f = record { _⟨$⟩_ = f - ; cong = P.cong f + ; cong = ≡.cong f } ------------------------------------------------------------------------ @@ -52,15 +52,15 @@ f ^ suc n = f ∘′ (f ^ n) ^-homo : ∀ f → Homomorphic₂ ℕ Endo _≡_ (f ^_) _+_ _∘′_ ^-homo f zero n = refl -^-homo f (suc m) n = P.cong (f ∘′_) (^-homo f m n) +^-homo f (suc m) n = ≡.cong (f ∘′_) (^-homo f m n) ------------------------------------------------------------------------ -- Structures ∘-isMagma : IsMagma _≡_ (Op₂ Endo ∋ _∘′_) ∘-isMagma = record - { isEquivalence = P.isEquivalence - ; ∙-cong = P.cong₂ _∘′_ + { isEquivalence = ≡.isEquivalence + ; ∙-cong = ≡.cong₂ _∘′_ } ∘-magma : Magma _ _ @@ -89,7 +89,7 @@ f ^ suc n = f ∘′ (f ^ n) ^-isSemigroupMorphism : ∀ f → IsSemigroupMorphism +-semigroup ∘-semigroup (f ^_) ^-isSemigroupMorphism f = record - { ⟦⟧-cong = P.cong (f ^_) + { ⟦⟧-cong = ≡.cong (f ^_) ; ∙-homo = ^-homo f } diff --git a/src/Function/Equality.agda b/src/Function/Equality.agda index 7451aaf3f7..647dca7416 100644 --- a/src/Function/Equality.agda +++ b/src/Function/Equality.agda @@ -21,8 +21,8 @@ open import Relation.Binary.Indexed.Heterogeneous using (IndexedSetoid; _=[_]⇒_) import Relation.Binary.Indexed.Heterogeneous.Construct.Trivial as Trivial -import Relation.Binary.PropositionalEquality.Core as P -import Relation.Binary.PropositionalEquality.Properties as P +import Relation.Binary.PropositionalEquality.Core as ≡ +import Relation.Binary.PropositionalEquality.Properties as ≡ ------------------------------------------------------------------------ -- Functions which preserve equality @@ -147,8 +147,8 @@ flip {B = B} f = record } →-to-⟶ : ∀ {a b ℓ} {A : Set a} {B : Setoid b ℓ} → - (A → Setoid.Carrier B) → P.setoid A ⟶ B + (A → Setoid.Carrier B) → ≡.setoid A ⟶ B →-to-⟶ {B = B} to = record { _⟨$⟩_ = to - ; cong = λ { P.refl → Setoid.refl B } + ; cong = λ { ≡.refl → Setoid.refl B } } diff --git a/src/Function/Related.agda b/src/Function/Related.agda index 9846b4d034..3bcb839237 100644 --- a/src/Function/Related.agda +++ b/src/Function/Related.agda @@ -27,8 +27,8 @@ open import Relation.Binary.Core using (_⇒_) open import Relation.Binary.Bundles using (Setoid; Preorder) open import Relation.Binary.Structures using (IsEquivalence; IsPreorder) open import Relation.Binary.Definitions using (Reflexive; Trans; Sym) -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 ≡ open import Data.Product.Base using (_,_; proj₁; proj₂; <_,_>) import Function.Related.Propositional as R @@ -90,12 +90,12 @@ infix 4 _∼[_]_ _∼[_]_ : ∀ {ℓ₁ ℓ₂} → Set ℓ₁ → Kind → Set ℓ₂ → Set _ A ∼[ implication ] B = A → B A ∼[ reverse-implication ] B = A ← B -A ∼[ equivalence ] B = Equivalence (P.setoid A) (P.setoid B) -A ∼[ injection ] B = Injection (P.setoid A) (P.setoid B) +A ∼[ equivalence ] B = Equivalence (≡.setoid A) (≡.setoid B) +A ∼[ injection ] B = Injection (≡.setoid A) (≡.setoid B) A ∼[ reverse-injection ] B = A ↢ B -A ∼[ left-inverse ] B = LeftInverse (P.setoid A) (P.setoid B) -A ∼[ surjection ] B = Surjection (P.setoid A) (P.setoid B) -A ∼[ bijection ] B = Inverse (P.setoid A) (P.setoid B) +A ∼[ left-inverse ] B = LeftInverse (≡.setoid A) (≡.setoid B) +A ∼[ surjection ] B = Surjection (≡.setoid A) (≡.setoid B) +A ∼[ bijection ] B = Inverse (≡.setoid A) (≡.setoid B) -- A non-infix synonym. @@ -125,7 +125,7 @@ fromRelated {K = left-inverse} record { to = to ; from = from ; left-inve B.mk↪ {to = to ⟨$⟩_} {from = from ⟨$⟩_} (strictlyInverseʳ⇒inverseʳ (to ⟨$⟩_) left-inverse-of) fromRelated {K = surjection} record { to = to ; surjective = surjective } with surjective ... | record { from = from ; right-inverse-of = right-inverse-of } = - B.mk↠ {to = to ⟨$⟩_} < from ⟨$⟩_ , (λ { x P.refl → right-inverse-of x }) > + B.mk↠ {to = to ⟨$⟩_} < from ⟨$⟩_ , (λ { x ≡.refl → right-inverse-of x }) > fromRelated {K = bijection} rel = B.mk↔ₛ′ (to ⟨$⟩_) (from ⟨$⟩_) right-inverse-of left-inverse-of where open Inverse rel @@ -145,7 +145,7 @@ fromRelated {K = bijection} rel = B.mk↔ₛ′ (to ⟨$⟩_) (from -- Actual equality also implies any kind of relatedness. ≡⇒ : ∀ {k ℓ} {X Y : Set ℓ} → X ≡ Y → X ∼[ k ] Y -≡⇒ P.refl = ↔⇒ Inv.id +≡⇒ ≡.refl = ↔⇒ Inv.id ------------------------------------------------------------------------ -- Special kinds of kinds @@ -313,7 +313,7 @@ K-refl {surjection} = Surj.id K-refl {bijection} = Inv.id K-reflexive : ∀ {k ℓ} → _≡_ ⇒ Related k {ℓ} -K-reflexive P.refl = K-refl +K-reflexive ≡.refl = K-refl K-trans : ∀ {k ℓ₁ ℓ₂ ℓ₃} → Trans (Related k {ℓ₁} {ℓ₂}) (Related k {ℓ₂} {ℓ₃}) @@ -383,7 +383,7 @@ module EquationalReasoning where _≡˘⟨_⟩_ : ∀ {k ℓ z} (X : Set ℓ) {Y : Set ℓ} {Z : Set z} → Y ≡ X → Y ∼[ k ] Z → X ∼[ k ] Z - X ≡˘⟨ Y≡X ⟩ Y⇔Z = X ∼⟨ ≡⇒ (P.sym Y≡X) ⟩ Y⇔Z + X ≡˘⟨ Y≡X ⟩ Y⇔Z = X ∼⟨ ≡⇒ (≡.sym Y≡X) ⟩ Y⇔Z _≡⟨_⟩_ : ∀ {k ℓ z} (X : Set ℓ) {Y : Set ℓ} {Z : Set z} → X ≡ Y → Y ∼[ k ] Z → X ∼[ k ] Z @@ -406,10 +406,10 @@ InducedPreorder₁ k S = record { _≈_ = _≡_ ; _≲_ = InducedRelation₁ k S ; isPreorder = record - { isEquivalence = P.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = reflexive ∘ K-reflexive ∘ - P.cong S + ≡.cong S ; trans = K-trans } } where open Preorder (K-preorder _ _) @@ -439,11 +439,11 @@ InducedPreorder₂ k _S_ = record { _≈_ = _≡_ ; _≲_ = InducedRelation₂ k _S_ ; isPreorder = record - { isEquivalence = P.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = λ x≡y {z} → reflexive $ K-reflexive $ - P.cong (_S_ z) x≡y + ≡.cong (_S_ z) x≡y ; trans = λ i↝j j↝k → K-trans i↝j j↝k } diff --git a/src/Function/Related/Propositional.agda b/src/Function/Related/Propositional.agda index d7f0610e78..589c36a6b3 100644 --- a/src/Function/Related/Propositional.agda +++ b/src/Function/Related/Propositional.agda @@ -13,8 +13,8 @@ open import Relation.Binary using (Rel; REL; Sym; Reflexive; Trans; IsEquivalence; Setoid; IsPreorder; Preorder) open import Function.Bundles open import Function.Base -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 ≡ open import Relation.Binary.Reasoning.Syntax open import Function.Properties.Surjection using (↠⇒↪; ↠⇒⇔) @@ -86,7 +86,7 @@ Related k A B = A ∼[ k ] B -- Propositional equality also implies any kind of relatedness. ≡⇒ : A ≡ B → A ∼[ k ] B -≡⇒ P.refl = ↔⇒ (Identity.↔-id _) +≡⇒ ≡.refl = ↔⇒ (Identity.↔-id _) ------------------------------------------------------------------------ -- Special kinds of kinds @@ -252,7 +252,7 @@ K-refl {k = surjection} = Identity.↠-id _ K-refl {k = bijection} = Identity.↔-id _ K-reflexive : _≡_ Relation.Binary.⇒ Related {a} k -K-reflexive P.refl = K-refl +K-reflexive ≡.refl = K-refl K-trans : Trans (Related {a} {b} k) (Related {b} {c} k) @@ -340,10 +340,10 @@ InducedPreorder₁ k P = record { _≈_ = _≡_ ; _≲_ = InducedRelation₁ k P ; isPreorder = record - { isEquivalence = P.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = reflexive ∘ K-reflexive ∘ - P.cong P + ≡.cong P ; trans = K-trans } } where open Preorder (K-preorder _ _) @@ -370,11 +370,11 @@ InducedPreorder₂ k _S_ = record { _≈_ = _≡_ ; _≲_ = InducedRelation₂ k _S_ ; isPreorder = record - { isEquivalence = P.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = λ x≡y {z} → reflexive $ K-reflexive $ - P.cong (_S_ z) x≡y + ≡.cong (_S_ z) x≡y ; trans = λ i↝j j↝k → K-trans i↝j j↝k } diff --git a/src/Function/Related/TypeIsomorphisms.agda b/src/Function/Related/TypeIsomorphisms.agda index d8a78e052c..48ed246bca 100644 --- a/src/Function/Related/TypeIsomorphisms.agda +++ b/src/Function/Related/TypeIsomorphisms.agda @@ -28,7 +28,7 @@ open import Function.Bundles open import Function.Related.Propositional import Function.Construct.Identity as Identity open import Relation.Binary hiding (_⇔_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) open import Relation.Nullary.Reflects using (invert) @@ -47,20 +47,20 @@ 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↔ₛ′ Product.assocʳ Product.assocˡ (λ _ → P.refl) (λ _ → P.refl) +Σ-assoc = mk↔ₛ′ Product.assocʳ Product.assocˡ (λ _ → ≡.refl) (λ _ → ≡.refl) -- × is commutative ×-comm : ∀ (A : Set a) (B : Set b) → (A × B) ↔ (B × A) -×-comm _ _ = mk↔ₛ′ Product.swap Product.swap (λ _ → P.refl) λ _ → P.refl +×-comm _ _ = mk↔ₛ′ Product.swap Product.swap (λ _ → ≡.refl) λ _ → ≡.refl -- × has ⊤ as its identity ×-identityˡ : ∀ ℓ → LeftIdentity {ℓ = ℓ} _↔_ ⊤ _×_ -×-identityˡ _ _ = mk↔ₛ′ proj₂ -,_ (λ _ → P.refl) (λ _ → P.refl) +×-identityˡ _ _ = mk↔ₛ′ proj₂ -,_ (λ _ → ≡.refl) (λ _ → ≡.refl) ×-identityʳ : ∀ ℓ → RightIdentity {ℓ = ℓ} _↔_ ⊤ _×_ -×-identityʳ _ _ = mk↔ₛ′ proj₁ (_, _) (λ _ → P.refl) (λ _ → P.refl) +×-identityʳ _ _ = mk↔ₛ′ proj₁ (_, _) (λ _ → ≡.refl) (λ _ → ≡.refl) ×-identity : ∀ ℓ → Identity _↔_ ⊤ _×_ ×-identity ℓ = ×-identityˡ ℓ , ×-identityʳ ℓ @@ -68,10 +68,10 @@ private -- × has ⊥ has its zero ×-zeroˡ : ∀ ℓ → LeftZero {ℓ = ℓ} _↔_ ⊥ _×_ -×-zeroˡ ℓ A = mk↔ₛ′ proj₁ < id , ⊥ₚ-elim > (λ _ → P.refl) (λ { () }) +×-zeroˡ ℓ A = mk↔ₛ′ proj₁ < id , ⊥ₚ-elim > (λ _ → ≡.refl) (λ { () }) ×-zeroʳ : ∀ ℓ → RightZero {ℓ = ℓ} _↔_ ⊥ _×_ -×-zeroʳ ℓ A = mk↔ₛ′ proj₂ < ⊥ₚ-elim , id > (λ _ → P.refl) (λ { () }) +×-zeroʳ ℓ A = mk↔ₛ′ proj₂ < ⊥ₚ-elim , id > (λ _ → ≡.refl) (λ { () }) ×-zero : ∀ ℓ → Zero _↔_ ⊥ _×_ ×-zero ℓ = ×-zeroˡ ℓ , ×-zeroʳ ℓ @@ -85,8 +85,8 @@ private ⊎-assoc ℓ _ _ _ = mk↔ₛ′ [ [ inj₁ , inj₂ ∘′ inj₁ ]′ , inj₂ ∘′ inj₂ ]′ [ inj₁ ∘′ inj₁ , [ inj₁ ∘′ inj₂ , inj₂ ]′ ]′ - [ (λ _ → P.refl) , [ (λ _ → P.refl) , (λ _ → P.refl) ] ] - [ [ (λ _ → P.refl) , (λ _ → P.refl) ] , (λ _ → P.refl) ] + [ (λ _ → ≡.refl) , [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] ] + [ [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] , (λ _ → ≡.refl) ] -- ⊎ is commutative @@ -96,12 +96,12 @@ private -- ⊎ has ⊥ as its identity ⊎-identityˡ : ∀ ℓ → LeftIdentity _↔_ (⊥ {ℓ}) _⊎_ -⊎-identityˡ _ _ = mk↔ₛ′ [ (λ ()) , id ]′ inj₂ (λ _ → P.refl) - [ (λ ()) , (λ _ → P.refl) ] +⊎-identityˡ _ _ = mk↔ₛ′ [ (λ ()) , id ]′ inj₂ (λ _ → ≡.refl) + [ (λ ()) , (λ _ → ≡.refl) ] ⊎-identityʳ : ∀ ℓ → RightIdentity _↔_ (⊥ {ℓ}) _⊎_ -⊎-identityʳ _ _ = mk↔ₛ′ [ id , (λ ()) ]′ inj₁ (λ _ → P.refl) - [ (λ _ → P.refl) , (λ ()) ] +⊎-identityʳ _ _ = mk↔ₛ′ [ id , (λ ()) ]′ inj₁ (λ _ → ≡.refl) + [ (λ _ → ≡.refl) , (λ ()) ] ⊎-identity : ∀ ℓ → Identity _↔_ ⊥ _⊎_ ⊎-identity ℓ = ⊎-identityˡ ℓ , ⊎-identityʳ ℓ @@ -115,15 +115,15 @@ private ×-distribˡ-⊎ ℓ _ _ _ = mk↔ₛ′ (uncurry λ x → [ inj₁ ∘′ (x ,_) , inj₂ ∘′ (x ,_) ]′) [ Product.map₂ inj₁ , Product.map₂ inj₂ ]′ - [ (λ _ → P.refl) , (λ _ → P.refl) ] - (uncurry λ _ → [ (λ _ → P.refl) , (λ _ → P.refl) ]) + [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] + (uncurry λ _ → [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ]) ×-distribʳ-⊎ : ∀ ℓ → _DistributesOverʳ_ {ℓ = ℓ} _↔_ _×_ _⊎_ ×-distribʳ-⊎ ℓ _ _ _ = mk↔ₛ′ (uncurry [ curry inj₁ , curry inj₂ ]′) [ Product.map₁ inj₁ , Product.map₁ inj₂ ]′ - [ (λ _ → P.refl) , (λ _ → P.refl) ] - (uncurry [ (λ _ _ → P.refl) , (λ _ _ → P.refl) ]) + [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] + (uncurry [ (λ _ _ → ≡.refl) , (λ _ _ → ≡.refl) ]) ×-distrib-⊎ : ∀ ℓ → _DistributesOver_ {ℓ = ℓ} _↔_ _×_ _⊎_ ×-distrib-⊎ ℓ = ×-distribˡ-⊎ ℓ , ×-distribʳ-⊎ ℓ @@ -244,11 +244,11 @@ private ΠΠ↔ΠΠ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) → ((x : A) (y : B) → P x y) ↔ ((y : B) (x : A) → P x y) -ΠΠ↔ΠΠ _ = mk↔ₛ′ flip flip (λ _ → P.refl) (λ _ → P.refl) +ΠΠ↔ΠΠ _ = mk↔ₛ′ flip flip (λ _ → ≡.refl) (λ _ → ≡.refl) ∃∃↔∃∃ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) → (∃₂ λ x y → P x y) ↔ (∃₂ λ y x → P x y) -∃∃↔∃∃ P = mk↔ₛ′ to from (λ _ → P.refl) (λ _ → P.refl) +∃∃↔∃∃ P = mk↔ₛ′ to from (λ _ → ≡.refl) (λ _ → ≡.refl) where to : (∃₂ λ x y → P x y) → (∃₂ λ y x → P x y) to (x , y , Pxy) = (y , x , Pxy) @@ -261,7 +261,7 @@ private Π↔Π : ∀ {A : Set a} {B : A → Set b} → ((x : A) → B x) ↔ ({x : A} → B x) -Π↔Π = mk↔ₛ′ _$- λ- (λ _ → P.refl) (λ _ → P.refl) +Π↔Π = mk↔ₛ′ _$- λ- (λ _ → ≡.refl) (λ _ → ≡.refl) ------------------------------------------------------------------------ -- _→_ preserves the symmetric relations @@ -280,11 +280,11 @@ private (λ f → from C↔D ∘ f ∘ to A↔B) (λ f → ext₂ λ x → begin to C↔D (from C↔D (f (to A↔B (from A↔B x)))) ≡⟨ strictlyInverseˡ C↔D _ ⟩ - f (to A↔B (from A↔B x)) ≡⟨ P.cong f $ strictlyInverseˡ A↔B x ⟩ + f (to A↔B (from A↔B x)) ≡⟨ ≡.cong f $ strictlyInverseˡ A↔B x ⟩ f x ∎) (λ f → ext₁ λ x → begin from C↔D (to C↔D (f (from A↔B (to A↔B x)))) ≡⟨ strictlyInverseʳ C↔D _ ⟩ - f (from A↔B (to A↔B x)) ≡⟨ P.cong f $ strictlyInverseʳ A↔B x ⟩ + f (from A↔B (to A↔B x)) ≡⟨ ≡.cong f $ strictlyInverseʳ A↔B x ⟩ f x ∎) where open Inverse; open ≡-Reasoning @@ -303,7 +303,7 @@ private ¬-cong : Extensionality a 0ℓ → Extensionality b 0ℓ → ∀ {k} {A : Set a} {B : Set b} → A ∼[ ⌊ k ⌋ ] B → (¬ A) ∼[ ⌊ k ⌋ ] (¬ B) -¬-cong extA extB A≈B = →-cong extA extB A≈B (K-reflexive P.refl) +¬-cong extA extB A≈B = →-cong extA extB A≈B (K-reflexive ≡.refl) ------------------------------------------------------------------------ -- _⇔_ preserves _⇔_ @@ -330,6 +330,6 @@ Related-cong {A = A} {B = B} {C = C} {D = D} A≈B C≈D = mk⇔ True↔ : ∀ {p} {P : Set p} (dec : Dec P) → ((p₁ p₂ : P) → p₁ ≡ p₂) → True dec ↔ P True↔ ( true because [p]) irr = - mk↔ₛ′ (λ _ → invert [p]) (λ _ → _) (irr _) (λ _ → P.refl) + mk↔ₛ′ (λ _ → invert [p]) (λ _ → _) (irr _) (λ _ → ≡.refl) True↔ (false because ofⁿ ¬p) _ = mk↔ₛ′ (λ()) (invert (ofⁿ ¬p)) (⊥-elim ∘ ¬p) (λ ()) diff --git a/src/Relation/Binary/Construct/Add/Infimum/NonStrict.agda b/src/Relation/Binary/Construct/Add/Infimum/NonStrict.agda index 5d5cb2a214..a151fe5065 100644 --- a/src/Relation/Binary/Construct/Add/Infimum/NonStrict.agda +++ b/src/Relation/Binary/Construct/Add/Infimum/NonStrict.agda @@ -20,9 +20,9 @@ module Relation.Binary.Construct.Add.Infimum.NonStrict open import Level using (_⊔_) open import Data.Sum.Base as Sum -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) -import Relation.Binary.PropositionalEquality.Properties as P +import Relation.Binary.PropositionalEquality.Properties as ≡ import Relation.Binary.Construct.Add.Infimum.Equality as Equality open import Relation.Nullary hiding (Irrelevant) open import Relation.Nullary.Construct.Add.Infimum @@ -60,8 +60,8 @@ data _≤₋_ : Rel (A ₋) (a ⊔ ℓ) where ≤₋-total ≤-total [ k ] [ l ] = Sum.map [_] [_] (≤-total k l) ≤₋-irrelevant : Irrelevant _≤_ → Irrelevant _≤₋_ -≤₋-irrelevant ≤-irr (⊥₋≤ k) (⊥₋≤ k) = P.refl -≤₋-irrelevant ≤-irr [ p ] [ q ] = P.cong _ (≤-irr p q) +≤₋-irrelevant ≤-irr (⊥₋≤ k) (⊥₋≤ k) = ≡.refl +≤₋-irrelevant ≤-irr [ p ] [ q ] = ≡.cong _ (≤-irr p q) ------------------------------------------------------------------------ -- Relational properties + propositional equality @@ -72,7 +72,7 @@ data _≤₋_ : Rel (A ₋) (a ⊔ ℓ) where ≤₋-antisym-≡ : Antisymmetric _≡_ _≤_ → Antisymmetric _≡_ _≤₋_ ≤₋-antisym-≡ antisym (⊥₋≤ _) (⊥₋≤ _) = refl -≤₋-antisym-≡ antisym [ p ] [ q ] = P.cong [_] (antisym p q) +≤₋-antisym-≡ antisym [ p ] [ q ] = ≡.cong [_] (antisym p q) ------------------------------------------------------------------------ -- Relational properties + setoid equality @@ -94,7 +94,7 @@ module _ {e} {_≈_ : Rel A e} where ≤₋-isPreorder-≡ : IsPreorder _≡_ _≤_ → IsPreorder _≡_ _≤₋_ ≤₋-isPreorder-≡ ≤-isPreorder = record - { isEquivalence = P.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = ≤₋-reflexive-≡ reflexive ; trans = ≤₋-trans trans } where open IsPreorder ≤-isPreorder diff --git a/src/Relation/Binary/Construct/Add/Infimum/Strict.agda b/src/Relation/Binary/Construct/Add/Infimum/Strict.agda index 5f3bdd7622..b9d47771db 100644 --- a/src/Relation/Binary/Construct/Add/Infimum/Strict.agda +++ b/src/Relation/Binary/Construct/Add/Infimum/Strict.agda @@ -21,9 +21,9 @@ module Relation.Binary.Construct.Add.Infimum.Strict open import Level using (_⊔_) open import Data.Product.Base using (_,_; map) open import Function.Base -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) -import Relation.Binary.PropositionalEquality.Properties as P +import Relation.Binary.PropositionalEquality.Properties as ≡ import Relation.Binary.Construct.Add.Infimum.Equality as Equality import Relation.Binary.Construct.Add.Infimum.NonStrict as NonStrict open import Relation.Nullary hiding (Irrelevant) @@ -59,8 +59,8 @@ data _<₋_ : Rel (A ₋) (a ⊔ ℓ) where <₋-dec _; _Respectsˡ_; _Respectsʳ_; _Respects_; _Respects₂_) open import Relation.Binary.Construct.Closure.Reflexive -open import Relation.Binary.PropositionalEquality.Core as PropEq using (_≡_; refl) -import Relation.Binary.PropositionalEquality.Properties as PropEq +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +import Relation.Binary.PropositionalEquality.Properties as ≡ open import Relation.Nullary import Relation.Nullary.Decidable as Dec open import Relation.Unary using (Pred) @@ -115,7 +115,7 @@ module _ {_~_ : Rel A ℓ} where isPreorder : Transitive _~_ → IsPreorder _≡_ _~ᵒ_ isPreorder ~-trans = record - { isEquivalence = PropEq.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = λ { refl → refl } ; trans = trans ~-trans } diff --git a/src/Relation/Binary/Construct/Closure/Transitive.agda b/src/Relation/Binary/Construct/Closure/Transitive.agda index e4154d49cd..3c91ce2318 100644 --- a/src/Relation/Binary/Construct/Closure/Transitive.agda +++ b/src/Relation/Binary/Construct/Closure/Transitive.agda @@ -15,7 +15,7 @@ open import Level open import Relation.Binary.Core using (Rel; _=[_]⇒_; _⇒_) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) private variable @@ -102,7 +102,7 @@ data Plus {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where module _ {_∼_ : Rel A ℓ} where []-injective : ∀ {x y p q} → (x [ _∼_ ]⁺ y ∋ [ p ]) ≡ [ q ] → p ≡ q - []-injective P.refl = P.refl + []-injective ≡.refl = ≡.refl -- See also ∼⁺⟨⟩-injectiveˡ and ∼⁺⟨⟩-injectiveʳ in -- Relation.Binary.Construct.Closure.Transitive.WithK. diff --git a/src/Relation/Binary/HeterogeneousEquality.agda b/src/Relation/Binary/HeterogeneousEquality.agda index 9181229eec..41d6b709a6 100644 --- a/src/Relation/Binary/HeterogeneousEquality.agda +++ b/src/Relation/Binary/HeterogeneousEquality.agda @@ -25,10 +25,10 @@ open import Relation.Binary.Indexed.Heterogeneous using (IndexedSetoid) open import Relation.Binary.Indexed.Heterogeneous.Construct.At using (_atₛ_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) open import Relation.Binary.Reasoning.Syntax -import Relation.Binary.PropositionalEquality.Properties as P +import Relation.Binary.PropositionalEquality.Properties as ≡ import Relation.Binary.HeterogeneousEquality.Core as Core private @@ -59,7 +59,7 @@ open Core public using (≅-to-≡; ≡-to-≅) ≅-to-type-≡ refl = refl ≅-to-subst-≡ : ∀ {A B : Set a} {x : A} {y : B} → (p : x ≅ y) → - P.subst (λ x → x) (≅-to-type-≡ p) x ≡ y + ≡.subst (λ x → x) (≅-to-type-≡ p) x ≡ y ≅-to-subst-≡ refl = refl ------------------------------------------------------------------------ @@ -89,7 +89,7 @@ subst₂-removable : ∀ (_∼_ : REL A B r) {x y u v} (eq₁ : x ≅ y) (eq₂ subst₂-removable _∼_ refl refl z = refl ≡-subst-removable : ∀ (P : Pred A p) {x y} (eq : x ≡ y) (z : P x) → - P.subst P eq z ≅ z + ≡.subst P eq z ≅ z ≡-subst-removable P refl z = refl cong : ∀ {A : Set a} {B : A → Set b} {x y} @@ -133,7 +133,7 @@ module _ {I : Set ℓ} (A : I → Set a) {B : {k : I} → A k → Set b} where icong-≡-subst-removable : {i j : I} (eq : i ≡ j) (f : {k : I} → (z : A k) → B z) (x : A i) → - f (P.subst A eq x) ≅ f x + f (≡.subst A eq x) ≅ f x icong-≡-subst-removable refl _ _ = refl ------------------------------------------------------------------------ @@ -183,13 +183,13 @@ indexedSetoid B = record } ≡↔≅ : ∀ {A : Set a} (B : A → Set b) {x : A} → - Inverse (P.setoid (B x)) ((indexedSetoid B) atₛ x) + Inverse (≡.setoid (B x)) ((indexedSetoid B) atₛ x) ≡↔≅ B = record { to = id ; to-cong = ≡-to-≅ ; from = id ; from-cong = ≅-to-≡ - ; inverse = (λ { P.refl → refl }) , λ { refl → P.refl } + ; inverse = (λ { ≡.refl → refl }) , λ { refl → ≡.refl } } decSetoid : Decidable {A = A} {B = A} (λ x y → x ≅ y) → @@ -211,7 +211,7 @@ isPreorder = record isPreorder-≡ : IsPreorder {A = A} _≡_ (λ x y → x ≅ y) isPreorder-≡ = record - { isEquivalence = P.isEquivalence + { isEquivalence = ≡.isEquivalence ; reflexive = reflexive ; trans = trans } diff --git a/src/Relation/Binary/Indexed/Heterogeneous/Bundles.agda b/src/Relation/Binary/Indexed/Heterogeneous/Bundles.agda index 29d4db5253..78811d8edd 100644 --- a/src/Relation/Binary/Indexed/Heterogeneous/Bundles.agda +++ b/src/Relation/Binary/Indexed/Heterogeneous/Bundles.agda @@ -14,7 +14,7 @@ module Relation.Binary.Indexed.Heterogeneous.Bundles where open import Function.Base open import Level using (suc; _⊔_) open import Relation.Binary.Core using (_⇒_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Indexed.Heterogeneous.Core open import Relation.Binary.Indexed.Heterogeneous.Structures diff --git a/src/Relation/Binary/Indexed/Heterogeneous/Core.agda b/src/Relation/Binary/Indexed/Heterogeneous/Core.agda index 7195bb9ec4..1be335bbfb 100644 --- a/src/Relation/Binary/Indexed/Heterogeneous/Core.agda +++ b/src/Relation/Binary/Indexed/Heterogeneous/Core.agda @@ -14,7 +14,7 @@ module Relation.Binary.Indexed.Heterogeneous.Core where open import Level import Relation.Binary.Core as B import Relation.Binary.Definitions as B -import Relation.Binary.PropositionalEquality.Core as P +import Relation.Binary.PropositionalEquality.Core as ≡ ------------------------------------------------------------------------ -- Indexed binary relations diff --git a/src/Relation/Binary/Indexed/Heterogeneous/Definitions.agda b/src/Relation/Binary/Indexed/Heterogeneous/Definitions.agda index f6657a6861..e14345b83c 100644 --- a/src/Relation/Binary/Indexed/Heterogeneous/Definitions.agda +++ b/src/Relation/Binary/Indexed/Heterogeneous/Definitions.agda @@ -14,7 +14,7 @@ module Relation.Binary.Indexed.Heterogeneous.Definitions where open import Level import Relation.Binary.Core as B import Relation.Binary.Definitions as B -import Relation.Binary.PropositionalEquality.Core as P +import Relation.Binary.PropositionalEquality.Core as ≡ open import Relation.Binary.Indexed.Heterogeneous.Core private diff --git a/src/Relation/Binary/Indexed/Heterogeneous/Structures.agda b/src/Relation/Binary/Indexed/Heterogeneous/Structures.agda index 4569ef2fb0..c084dadc8c 100644 --- a/src/Relation/Binary/Indexed/Heterogeneous/Structures.agda +++ b/src/Relation/Binary/Indexed/Heterogeneous/Structures.agda @@ -18,7 +18,7 @@ module Relation.Binary.Indexed.Heterogeneous.Structures open import Function.Base open import Level using (suc; _⊔_) open import Relation.Binary.Core using (_⇒_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Indexed.Heterogeneous.Definitions ------------------------------------------------------------------------ @@ -31,7 +31,7 @@ record IsIndexedEquivalence : Set (i ⊔ a ⊔ ℓ) where trans : Transitive A _≈_ reflexive : ∀ {i} → _≡_ ⟨ _⇒_ ⟩ _≈_ {i} - reflexive P.refl = refl + reflexive ≡.refl = refl record IsIndexedPreorder {ℓ₂} (_≲_ : IRel A ℓ₂) : Set (i ⊔ a ⊔ ℓ ⊔ ℓ₂) where diff --git a/src/Relation/Binary/Indexed/Homogeneous/Bundles.agda b/src/Relation/Binary/Indexed/Homogeneous/Bundles.agda index 4bbd118c13..4fccb7fca9 100644 --- a/src/Relation/Binary/Indexed/Homogeneous/Bundles.agda +++ b/src/Relation/Binary/Indexed/Homogeneous/Bundles.agda @@ -16,7 +16,7 @@ open import Function.Base using (_⟨_⟩_) open import Level using (Level; _⊔_; suc) open import Relation.Binary.Core using (_⇒_; Rel) open import Relation.Binary.Bundles as B -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary.Negation using (¬_) open import Relation.Binary.Indexed.Homogeneous.Core open import Relation.Binary.Indexed.Homogeneous.Structures diff --git a/src/Relation/Binary/Indexed/Homogeneous/Structures.agda b/src/Relation/Binary/Indexed/Homogeneous/Structures.agda index 77438dafdd..f5a941b65b 100644 --- a/src/Relation/Binary/Indexed/Homogeneous/Structures.agda +++ b/src/Relation/Binary/Indexed/Homogeneous/Structures.agda @@ -23,7 +23,7 @@ open import Level using (Level; _⊔_; suc) open import Relation.Binary.Core using (_⇒_) import Relation.Binary.Definitions as B import Relation.Binary.Structures as B -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Indexed.Homogeneous.Definitions ------------------------------------------------------------------------ @@ -40,12 +40,12 @@ record IsIndexedEquivalence : Set (i ⊔ a ⊔ ℓ) where transᵢ : Transitive A _≈ᵢ_ reflexiveᵢ : ∀ {i} → _≡_ ⟨ _⇒_ ⟩ _≈ᵢ_ {i} - reflexiveᵢ P.refl = reflᵢ + reflexiveᵢ ≡.refl = reflᵢ -- Lift properties reflexive : _≡_ ⇒ (Lift A _≈ᵢ_) - reflexive P.refl i = reflᵢ + reflexive ≡.refl i = reflᵢ refl : B.Reflexive (Lift A _≈ᵢ_) refl i = reflᵢ diff --git a/src/Relation/Binary/Properties/Setoid.agda b/src/Relation/Binary/Properties/Setoid.agda index 968256077f..f21eb2f4df 100644 --- a/src/Relation/Binary/Properties/Setoid.agda +++ b/src/Relation/Binary/Properties/Setoid.agda @@ -9,7 +9,7 @@ open import Data.Product.Base using (_,_) open import Function.Base using (_∘_; id; _$_; flip) open import Relation.Nullary.Negation.Core using (¬_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Bundles using (Setoid; Preorder; Poset) open import Relation.Binary.Definitions using (Symmetric; _Respectsˡ_; _Respectsʳ_; _Respects₂_) @@ -26,9 +26,9 @@ open Setoid S isPreorder : IsPreorder _≡_ _≈_ isPreorder = record { isEquivalence = record - { refl = P.refl - ; sym = P.sym - ; trans = P.trans + { refl = ≡.refl + ; sym = ≡.sym + ; trans = ≡.trans } ; reflexive = reflexive ; trans = trans diff --git a/src/Tactic/Cong.agda b/src/Tactic/Cong.agda index 846fde2484..bf8a743306 100644 --- a/src/Tactic/Cong.agda +++ b/src/Tactic/Cong.agda @@ -36,7 +36,7 @@ open import Data.Unit.Base using (⊤) open import Data.Word.Base as Word using (toℕ) open import Data.Product.Base using (_×_; map₁; _,_) -open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_; refl; cong) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl; cong) -- 'Data.String.Properties' defines this via 'Dec', so let's use the -- builtin for maximum speed. From 85e1589fb40c06d9db0fcb160522e4b72e39cd03 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 13 Feb 2024 09:43:34 +0000 Subject: [PATCH 3/5] Qualified import of `HeterogeneousEquality.Core` fixing #2280 --- src/Algebra/Construct/LiftedChoice.agda | 6 +-- src/Algebra/Operations/CommutativeMonoid.agda | 6 +-- .../Properties/CommutativeMonoid/Sum.agda | 6 +-- src/Algebra/Properties/Monoid/Mult.agda | 4 +- .../Properties/Monoid/Mult/TCOptimised.agda | 4 +- src/Algebra/Properties/Semiring/Exp.agda | 2 +- src/Data/Container/Indexed/WithK.agda | 14 +++--- .../Container/Relation/Binary/Pointwise.agda | 2 +- .../Relation/Binary/Pointwise/Properties.agda | 10 ++-- src/Data/Digit.agda | 6 +-- src/Data/Fin/Substitution/Example.agda | 2 +- src/Data/Graph/Acyclic.agda | 18 ++++---- .../Propositional/Properties/Core.agda | 10 ++-- src/Data/List/Relation/Binary/Pointwise.agda | 26 +++++------ .../Relation/Binary/Pointwise/Properties.agda | 6 +-- .../Prefix/Heterogeneous/Properties.agda | 8 ++-- .../Binary/Sublist/Heterogeneous.agda | 2 +- .../Sublist/Heterogeneous/Properties.agda | 46 +++++++++---------- .../List/Relation/Binary/Sublist/Setoid.agda | 2 +- .../Suffix/Heterogeneous/Properties.agda | 6 +-- .../List/Relation/Ternary/Interleaving.agda | 6 +-- src/Data/List/Relation/Unary/All.agda | 10 ++-- .../List/Relation/Unary/Any/Properties.agda | 46 +++++++++---------- .../Unary/Enumerates/Setoid/Properties.agda | 4 +- src/Data/Maybe/Relation/Binary/Connected.agda | 4 +- src/Data/Maybe/Relation/Binary/Pointwise.agda | 8 ++-- src/Data/Maybe/Relation/Unary/All.agda | 4 +- src/Data/Maybe/Relation/Unary/Any.agda | 2 +- src/Data/Nat/Coprimality.agda | 4 +- src/Data/Nat/GCD.agda | 28 +++++------ src/Data/Nat/LCM.agda | 6 +-- .../Product/Function/Dependent/Setoid.agda | 22 ++++----- .../Relation/Binary/Pointwise/Dependent.agda | 2 +- .../Binary/Pointwise/Dependent/WithK.agda | 14 +++--- .../Binary/Pointwise/NonDependent.agda | 10 ++-- src/Data/Sum/Relation/Binary/LeftOrder.agda | 2 +- src/Data/Sum/Relation/Binary/Pointwise.agda | 12 ++--- .../Relation/Binary/Equality/Setoid.agda | 4 +- src/Data/Vec/Properties/WithK.agda | 8 ++-- src/Data/Vec/Recursive/Properties.agda | 32 ++++++------- src/Data/Vec/Relation/Binary/Lex/Core.agda | 8 ++-- src/Data/Vec/Relation/Binary/Lex/Strict.agda | 10 ++-- .../Binary/Pointwise/Extensional.agda | 8 ++-- .../Relation/Binary/Pointwise/Inductive.agda | 12 ++--- src/Data/Vec/Relation/Unary/All.agda | 8 ++-- .../Vec/Relation/Unary/Any/Properties.agda | 30 ++++++------ .../Unary/Unique/Setoid/Properties.agda | 6 +-- 47 files changed, 248 insertions(+), 248 deletions(-) diff --git a/src/Algebra/Construct/LiftedChoice.agda b/src/Algebra/Construct/LiftedChoice.agda index 9b141f8f7e..2377650cf8 100644 --- a/src/Algebra/Construct/LiftedChoice.agda +++ b/src/Algebra/Construct/LiftedChoice.agda @@ -17,7 +17,7 @@ open import Relation.Nullary using (¬_; yes; no) open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]) open import Data.Product.Base using (_×_; _,_) open import Level using (Level; _⊔_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Unary using (Pred) import Relation.Binary.Reasoning.Setoid as ≈-Reasoning @@ -55,8 +55,8 @@ module _ {_≈_ : Rel B ℓ} {_∙_ : Op₂ B} sel-≡ : Selective _≡_ _◦_ sel-≡ x y with M.sel (f x) (f y) - ... | inj₁ _ = inj₁ P.refl - ... | inj₂ _ = inj₂ P.refl + ... | inj₁ _ = inj₁ ≡.refl + ... | inj₂ _ = inj₂ ≡.refl distrib : ∀ x y → ((f x) ∙ (f y)) ≈ f (x ◦ y) distrib x y with M.sel (f x) (f y) diff --git a/src/Algebra/Operations/CommutativeMonoid.agda b/src/Algebra/Operations/CommutativeMonoid.agda index 52d72af156..626d7570fb 100644 --- a/src/Algebra/Operations/CommutativeMonoid.agda +++ b/src/Algebra/Operations/CommutativeMonoid.agda @@ -12,7 +12,7 @@ open import Data.Fin.Base using (Fin; zero) open import Data.Nat.Base as ℕ using (ℕ; zero; suc) open import Function.Base using (_∘_) open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) module Algebra.Operations.CommutativeMonoid {s₁ s₂} (CM : CommutativeMonoid s₁ s₂) @@ -58,7 +58,7 @@ suc n ×′ x = x + n ×′ x ×-congʳ (suc n) x≈x′ = +-cong x≈x′ (×-congʳ n x≈x′) ×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_ -×-cong {u} P.refl x≈x′ = ×-congʳ u x≈x′ +×-cong {u} ≡.refl x≈x′ = ×-congʳ u x≈x′ -- _×_ is homomorphic with respect to _ℕ+_/_+_. @@ -98,7 +98,7 @@ suc n ×′ x = x + n ×′ x -- _×′_ preserves equality. ×′-cong : _×′_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_ -×′-cong {n} {_} {x} {y} P.refl x≈y = begin +×′-cong {n} {_} {x} {y} ≡.refl x≈y = begin n ×′ x ≈⟨ sym (×≈×′ n x) ⟩ n × x ≈⟨ ×-congʳ n x≈y ⟩ n × y ≈⟨ ×≈×′ n y ⟩ diff --git a/src/Algebra/Properties/CommutativeMonoid/Sum.agda b/src/Algebra/Properties/CommutativeMonoid/Sum.agda index eed54de4ed..351363c0d6 100644 --- a/src/Algebra/Properties/CommutativeMonoid/Sum.agda +++ b/src/Algebra/Properties/CommutativeMonoid/Sum.agda @@ -14,7 +14,7 @@ open import Data.Fin.Permutation as Perm using (Permutation; _⟨$⟩ˡ_; _⟨$ open import Data.Fin.Patterns using (0F) open import Data.Vec.Functional open import Function.Base using (_∘_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary.Negation using (contradiction) module Algebra.Properties.CommutativeMonoid.Sum @@ -90,9 +90,9 @@ sum-permute {zero} {suc n} f π = contradiction π (Perm.refute λ()) sum-permute {suc m} {zero} f π = contradiction π (Perm.refute λ()) sum-permute {suc m} {suc n} f π = begin sum f ≡⟨⟩ - f 0F + sum f/0 ≡⟨ P.cong (_+ sum f/0) (P.cong f (Perm.inverseʳ π)) ⟨ + f 0F + sum f/0 ≡⟨ ≡.cong (_+ sum f/0) (≡.cong f (Perm.inverseʳ π)) ⟨ πf π₀ + sum f/0 ≈⟨ +-congˡ (sum-permute f/0 (Perm.remove π₀ π)) ⟩ - πf π₀ + sum (rearrange (π/0 ⟨$⟩ʳ_) f/0) ≡⟨ P.cong (πf π₀ +_) (sum-cong-≗ (P.cong f ∘ Perm.punchIn-permute′ π 0F)) ⟨ + πf π₀ + sum (rearrange (π/0 ⟨$⟩ʳ_) f/0) ≡⟨ ≡.cong (πf π₀ +_) (sum-cong-≗ (≡.cong f ∘ Perm.punchIn-permute′ π 0F)) ⟨ πf π₀ + sum (removeAt πf π₀) ≈⟨ sym (sum-remove πf) ⟩ sum πf ∎ where diff --git a/src/Algebra/Properties/Monoid/Mult.agda b/src/Algebra/Properties/Monoid/Mult.agda index 2963523b77..ce837c8134 100644 --- a/src/Algebra/Properties/Monoid/Mult.agda +++ b/src/Algebra/Properties/Monoid/Mult.agda @@ -9,7 +9,7 @@ open import Algebra.Bundles using (Monoid) open import Data.Nat.Base as ℕ using (ℕ; zero; suc; NonZero) open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) module Algebra.Properties.Monoid.Mult {a ℓ} (M : Monoid a ℓ) where @@ -44,7 +44,7 @@ open import Algebra.Definitions.RawMonoid rawMonoid public ×-congʳ (suc n) x≈x′ = +-cong x≈x′ (×-congʳ n x≈x′) ×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_ -×-cong {n} P.refl x≈x′ = ×-congʳ n x≈x′ +×-cong {n} ≡.refl x≈x′ = ×-congʳ n x≈x′ ×-congˡ : ∀ {x} → (_× x) Preserves _≡_ ⟶ _≈_ ×-congˡ m≡n = ×-cong m≡n refl diff --git a/src/Algebra/Properties/Monoid/Mult/TCOptimised.agda b/src/Algebra/Properties/Monoid/Mult/TCOptimised.agda index c862c5e3f2..a147954cb0 100644 --- a/src/Algebra/Properties/Monoid/Mult/TCOptimised.agda +++ b/src/Algebra/Properties/Monoid/Mult/TCOptimised.agda @@ -10,7 +10,7 @@ open import Algebra.Bundles using (Monoid) open import Data.Nat.Base as ℕ using (ℕ; zero; suc) open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) module Algebra.Properties.Monoid.Mult.TCOptimised {a ℓ} (M : Monoid a ℓ) where @@ -75,7 +75,7 @@ open import Algebra.Definitions.RawMonoid rawMonoid public ×-congʳ (suc n@(suc _)) x≈y = +-cong (×-congʳ n x≈y) x≈y ×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_ -×-cong {n} P.refl x≈y = ×-congʳ n x≈y +×-cong {n} ≡.refl x≈y = ×-congʳ n x≈y ×-assocˡ : ∀ x m n → m × (n × x) ≈ (m ℕ.* n) × x ×-assocˡ x m n = begin diff --git a/src/Algebra/Properties/Semiring/Exp.agda b/src/Algebra/Properties/Semiring/Exp.agda index ed8b094520..f4adeee842 100644 --- a/src/Algebra/Properties/Semiring/Exp.agda +++ b/src/Algebra/Properties/Semiring/Exp.agda @@ -9,7 +9,7 @@ open import Algebra open import Data.Nat.Base as ℕ using (ℕ; zero; suc) open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) import Data.Nat.Properties as ℕ module Algebra.Properties.Semiring.Exp diff --git a/src/Data/Container/Indexed/WithK.agda b/src/Data/Container/Indexed/WithK.agda index 9897402264..bda5a15e34 100644 --- a/src/Data/Container/Indexed/WithK.agda +++ b/src/Data/Container/Indexed/WithK.agda @@ -20,8 +20,8 @@ open import Data.Product.Base open import Function.Base renaming (id to ⟨id⟩; _∘_ to _⟨∘⟩_) open import Level open import Relation.Unary using (Pred; _⊆_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; refl) -open import Relation.Binary.HeterogeneousEquality as H using (_≅_; refl) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +open import Relation.Binary.HeterogeneousEquality as ≅ using (_≅_; refl) open import Relation.Binary.Indexed.Heterogeneous ------------------------------------------------------------------------ @@ -43,7 +43,7 @@ private {xs : ⟦ C ⟧ X o₁} {ys : ⟦ C ⟧ X o₂} → Extensionality r ℓ → Eq C X X (λ x₁ x₂ → x₁ ≅ x₂) xs ys → xs ≅ ys Eq⇒≅ {xs = c , k} {.c , k′} ext (refl , refl , k≈k′) = - H.cong (_,_ c) (ext (λ _ → refl) (λ r → k≈k′ r r refl)) + ≅.cong (_,_ c) (ext (λ _ → refl) (λ r → k≈k′ r r refl)) setoid : ∀ {i o c r s} {I : Set i} {O : Set o} → Container I O c r → IndexedSetoid I s _ → IndexedSetoid O _ _ @@ -122,7 +122,7 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where module Y = IndexedSetoid Y lemma : ∀ {i j} (eq : i ≡ j) {x} → - P.subst Y.Carrier eq (f x) Y.≈ f (P.subst X eq x) + ≡.subst Y.Carrier eq (f x) Y.≈ f (≡.subst X eq x) lemma refl = Y.refl -- In fact, all natural functions of the right type are container @@ -135,7 +135,7 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where Eq C₂ X.Carrier X.Carrier X._≈_ (proj₁ nt X.Carrier xs) (⟪ m ⟫ X.Carrier {o} xs) complete {C₁} {C₂} (nt , nat) = m , (λ X xs → nat X - (λ { (r , eq) → P.subst (IndexedSetoid.Carrier X) eq (proj₂ xs r) }) + (λ { (r , eq) → ≡.subst (IndexedSetoid.Carrier X) eq (proj₂ xs r) }) (proj₁ xs , (λ r → r , refl))) where @@ -167,9 +167,9 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where module X = IndexedSetoid X lemma : ∀ {i j k} (eq₁ : i ≡ j) (eq₂ : j ≡ k) {x} → - P.subst X.Carrier (P.trans eq₁ eq₂) x + ≡.subst X.Carrier (≡.trans eq₁ eq₂) x X.≈ - P.subst X.Carrier eq₂ (P.subst X.Carrier eq₁ x) + ≡.subst X.Carrier eq₂ (≡.subst X.Carrier eq₁ x) lemma refl refl = X.refl ------------------------------------------------------------------------ diff --git a/src/Data/Container/Relation/Binary/Pointwise.agda b/src/Data/Container/Relation/Binary/Pointwise.agda index 2bd20bc1c5..bc66fe077d 100644 --- a/src/Data/Container/Relation/Binary/Pointwise.agda +++ b/src/Data/Container/Relation/Binary/Pointwise.agda @@ -12,7 +12,7 @@ open import Data.Product.Base using (_,_; Σ-syntax; -,_; proj₁; proj₂) open import Function.Base using (_∘_) open import Level using (_⊔_) open import Relation.Binary.Core using (REL; _⇒_) -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; subst; cong) open import Data.Container.Core using (Container; ⟦_⟧) diff --git a/src/Data/Container/Relation/Binary/Pointwise/Properties.agda b/src/Data/Container/Relation/Binary/Pointwise/Properties.agda index 45b824cc56..978bfc55e4 100644 --- a/src/Data/Container/Relation/Binary/Pointwise/Properties.agda +++ b/src/Data/Container/Relation/Binary/Pointwise/Properties.agda @@ -17,19 +17,19 @@ open import Relation.Binary.Core using (Rel) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive) open import Relation.Binary.Core using (Rel) -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; subst; cong) module _ {s p x r} {X : Set x} (C : Container s p) (R : Rel X r) where refl : Reflexive R → Reflexive (Pointwise C R) - refl R-refl = P.refl , λ p → R-refl + refl R-refl = ≡.refl , λ p → R-refl sym : Symmetric R → Symmetric (Pointwise C R) - sym R-sym (P.refl , f) = P.refl , λ p → R-sym (f p) + sym R-sym (≡.refl , f) = ≡.refl , λ p → R-sym (f p) trans : Transitive R → Transitive (Pointwise C R) - trans R-trans (P.refl , f) (P.refl , g) = P.refl , λ p → R-trans (f p) (g p) + trans R-trans (≡.refl , f) (≡.refl , g) = ≡.refl , λ p → R-trans (f p) (g p) private @@ -38,4 +38,4 @@ private Eq⇒≡ : ∀ {s p x} {C : Container s p} {X : Set x} {xs ys : ⟦ C ⟧ X} → Extensionality p x → Pointwise C _≡_ xs ys → xs ≡ ys - Eq⇒≡ ext (P.refl , f≈f′) = cong -,_ (ext f≈f′) + Eq⇒≡ ext (≡.refl , f≈f′) = cong -,_ (ext f≈f′) diff --git a/src/Data/Digit.agda b/src/Data/Digit.agda index 207b0280b7..82582aa472 100644 --- a/src/Data/Digit.agda +++ b/src/Data/Digit.agda @@ -21,7 +21,7 @@ open import Data.Nat.DivMod open import Data.Nat.Induction open import Relation.Nullary.Decidable using (True; does; toWitness) open import Relation.Binary.Definitions using (Decidable) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) open import Function.Base using (_$_) ------------------------------------------------------------------------ @@ -85,7 +85,7 @@ toDigits base@(suc (suc k)) n = <′-rec Pred helper n Pred = λ n → ∃ λ ds → fromDigits ds ≡ n cons : ∀ {m} (r : Digit base) → Pred m → Pred (toℕ r + m * base) - cons r (ds , eq) = (r ∷ ds , P.cong (λ i → toℕ r + i * base) eq) + cons r (ds , eq) = (r ∷ ds , ≡.cong (λ i → toℕ r + i * base) eq) open ≤-Reasoning open +-*-Solver @@ -104,7 +104,7 @@ toDigits base@(suc (suc k)) n = <′-rec Pred helper n helper : ∀ n → <′-Rec Pred n → Pred n helper n rec with n divMod base - ... | result zero r eq = ([ r ] , P.sym eq) + ... | result zero r eq = ([ r ] , ≡.sym eq) ... | result (suc x) r refl = cons r (rec (lem x k (toℕ r))) ------------------------------------------------------------------------ diff --git a/src/Data/Fin/Substitution/Example.agda b/src/Data/Fin/Substitution/Example.agda index 8b00a76843..5e5205ae14 100644 --- a/src/Data/Fin/Substitution/Example.agda +++ b/src/Data/Fin/Substitution/Example.agda @@ -22,7 +22,7 @@ open import Data.Fin.Base using (Fin) open import Data.Vec.Base open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl; sym; cong; cong₂) -open PropEq.≡-Reasoning +open ≡.≡-Reasoning open import Relation.Binary.Construct.Closure.ReflexiveTransitive using (Star; ε; _◅_) diff --git a/src/Data/Graph/Acyclic.agda b/src/Data/Graph/Acyclic.agda index c842cf3d30..26af8a3fd9 100644 --- a/src/Data/Graph/Acyclic.agda +++ b/src/Data/Graph/Acyclic.agda @@ -27,7 +27,7 @@ open import Data.Vec.Base as Vec using (Vec; []; _∷_) open import Data.List.Base as List using (List; []; _∷_) open import Function.Base using (_$_; _∘′_; _∘_; id) open import Relation.Nullary -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) ------------------------------------------------------------------------ -- A lemma @@ -186,7 +186,7 @@ private test-nodes : nodes example ≡ (# 0 , 0) ∷ (# 1 , 1) ∷ (# 2 , 2) ∷ (# 3 , 3) ∷ (# 4 , 4) ∷ [] - test-nodes = P.refl + test-nodes = ≡.refl module _ {ℓ e} {N : Set ℓ} {E : Set e} where @@ -213,7 +213,7 @@ private test-edges : edges example ≡ (# 1 , 10 , # 1) ∷ (# 1 , 11 , # 1) ∷ (# 2 , 12 , # 0) ∷ [] - test-edges = P.refl + test-edges = ≡.refl -- The successors of a given node i (edge label × node number relative -- to i). @@ -225,7 +225,7 @@ sucs g i = successors $ head (g [ i ]) private test-sucs : sucs example (# 1) ≡ (10 , # 1) ∷ (11 , # 1) ∷ [] - test-sucs = P.refl + test-sucs = ≡.refl -- The predecessors of a given node i (node number relative to i × -- edge label). @@ -238,13 +238,13 @@ preds (c & g) (suc i) = (List.map (Prod.map suc id) $ preds g i) where p : ∀ {e} {E : Set e} {n} (i : Fin n) → E × Fin n → Maybe (Fin′ (suc i) × E) - p i (e , j) = Maybe.map (λ{ P.refl → zero , e }) (decToMaybe (i ≟ j)) + p i (e , j) = Maybe.map (λ{ ≡.refl → zero , e }) (decToMaybe (i ≟ j)) private test-preds : preds example (# 3) ≡ (# 1 , 10) ∷ (# 1 , 11) ∷ (# 2 , 12) ∷ [] - test-preds = P.refl + test-preds = ≡.refl ------------------------------------------------------------------------ -- Operations @@ -272,7 +272,7 @@ private context (# 3 , 3) [] & context (# 4 , 4) [] & ∅) - test-number = P.refl + test-number = ≡.refl -- Reverses all the edges in the graph. @@ -290,7 +290,7 @@ reverse {N = N} {E} g = private test-reverse : reverse (reverse example) ≡ example - test-reverse = P.refl + test-reverse = ≡.refl ------------------------------------------------------------------------ -- Views @@ -330,4 +330,4 @@ private node 3 [] ∷ node 4 [] ∷ [] - test-toForest = P.refl + test-toForest = ≡.refl diff --git a/src/Data/List/Membership/Propositional/Properties/Core.agda b/src/Data/List/Membership/Propositional/Properties/Core.agda index f7543d05f8..568b567b37 100644 --- a/src/Data/List/Membership/Propositional/Properties/Core.agda +++ b/src/Data/List/Membership/Propositional/Properties/Core.agda @@ -20,7 +20,7 @@ open import Data.List.Membership.Propositional open import Data.Product.Base as Product using (_,_; proj₁; proj₂; uncurry′; ∃; _×_) open import Level using (Level) -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) open import Relation.Unary using (Pred; _⊆_) @@ -37,8 +37,8 @@ map∘find : ∀ {P : Pred A p} {xs} {f : _≡_ (proj₁ p′) ⊆ P} → f refl ≡ proj₂ (proj₂ p′) → Any.map f (proj₁ (proj₂ p′)) ≡ p -map∘find (here p) hyp = P.cong here hyp -map∘find (there p) hyp = P.cong there (map∘find p hyp) +map∘find (here p) hyp = ≡.cong here hyp +map∘find (there p) hyp = ≡.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) → @@ -61,13 +61,13 @@ find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl lose∘find : ∀ {P : Pred A p} {xs : List A} (p : Any P xs) → uncurry′ lose (proj₂ (find p)) ≡ p -lose∘find p = map∘find p P.refl +lose∘find p = map∘find p ≡.refl find∘lose : ∀ (P : Pred A p) {x xs} (x∈xs : x ∈ xs) (pp : P x) → find {P = P} (lose x∈xs pp) ≡ (x , x∈xs , pp) find∘lose P x∈xs p - rewrite find∘map x∈xs (flip (P.subst P) p) + rewrite find∘map x∈xs (flip (≡.subst P) p) | find-∈ x∈xs = refl diff --git a/src/Data/List/Relation/Binary/Pointwise.agda b/src/Data/List/Relation/Binary/Pointwise.agda index 2728418ddc..76e9d81f54 100644 --- a/src/Data/List/Relation/Binary/Pointwise.agda +++ b/src/Data/List/Relation/Binary/Pointwise.agda @@ -29,8 +29,8 @@ open import Relation.Binary.Core renaming (Rel to Rel₂) open import Relation.Binary.Definitions using (_Respects_; _Respects₂_) open import Relation.Binary.Bundles using (Setoid; DecSetoid; Preorder; Poset) open import Relation.Binary.Structures using (IsEquivalence; IsDecEquivalence; IsPartialOrder; IsPreorder) -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 @@ -129,8 +129,8 @@ AllPairs-resp-Pointwise resp@(respₗ , respᵣ) (x∼y ∷ xs) (px ∷ pxs) = -- length Pointwise-length : Pointwise R xs ys → length xs ≡ length ys -Pointwise-length [] = P.refl -Pointwise-length (x∼y ∷ xs∼ys) = P.cong ℕ.suc (Pointwise-length xs∼ys) +Pointwise-length [] = ≡.refl +Pointwise-length (x∼y ∷ xs∼ys) = ≡.cong ℕ.suc (Pointwise-length xs∼ys) ------------------------------------------------------------------------ -- tabulate @@ -162,9 +162,9 @@ tabulate⁻ {n = suc n} (x∼y ∷ xs∼ys) (fsuc i) = tabulate⁻ xs∼ys i ++-cancelʳ (y ∷ ys) (z ∷ zs) (y∼z ∷ ys∼zs) = y∼z ∷ (++-cancelʳ ys zs ys∼zs) -- Impossible cases ++-cancelʳ {xs = xs} [] (z ∷ zs) eq = - contradiction (P.trans (Pointwise-length eq) (length-++ (z ∷ zs))) (m≢1+n+m (length xs)) + contradiction (≡.trans (Pointwise-length eq) (length-++ (z ∷ zs))) (m≢1+n+m (length xs)) ++-cancelʳ {xs = xs} (y ∷ ys) [] eq = - contradiction (P.trans (P.sym (length-++ (y ∷ ys))) (Pointwise-length eq)) (m≢1+n+m (length xs) ∘ P.sym) + contradiction (≡.trans (≡.sym (length-++ (y ∷ ys))) (Pointwise-length eq)) (m≢1+n+m (length xs) ∘ ≡.sym) ------------------------------------------------------------------------ -- concat @@ -246,8 +246,8 @@ lookup⁻ : length xs ≡ length ys → (∀ {i j} → toℕ i ≡ toℕ j → R (lookup xs i) (lookup ys j)) → Pointwise R xs ys lookup⁻ {xs = []} {ys = []} _ _ = [] -lookup⁻ {xs = _ ∷ _} {ys = _ ∷ _} |xs|≡|ys| eq = eq {fzero} P.refl ∷ - lookup⁻ (suc-injective |xs|≡|ys|) (eq ∘ P.cong ℕ.suc) +lookup⁻ {xs = _ ∷ _} {ys = _ ∷ _} |xs|≡|ys| eq = eq {fzero} ≡.refl ∷ + lookup⁻ (suc-injective |xs|≡|ys|) (eq ∘ ≡.cong ℕ.suc) lookup⁺ : ∀ (Rxys : Pointwise R xs ys) → ∀ i → (let j = cast (Pointwise-length Rxys) i) → @@ -260,14 +260,14 @@ lookup⁺ (_ ∷ Rxys) (fsuc i) = lookup⁺ Rxys i ------------------------------------------------------------------------ Pointwise-≡⇒≡ : Pointwise {A = A} _≡_ ⇒ _≡_ -Pointwise-≡⇒≡ [] = P.refl -Pointwise-≡⇒≡ (P.refl ∷ xs∼ys) with Pointwise-≡⇒≡ xs∼ys -... | P.refl = P.refl +Pointwise-≡⇒≡ [] = ≡.refl +Pointwise-≡⇒≡ (≡.refl ∷ xs∼ys) with Pointwise-≡⇒≡ xs∼ys +... | ≡.refl = ≡.refl ≡⇒Pointwise-≡ : _≡_ ⇒ Pointwise {A = A} _≡_ -≡⇒Pointwise-≡ P.refl = refl P.refl +≡⇒Pointwise-≡ ≡.refl = refl ≡.refl -Pointwise-≡↔≡ : Inverse (setoid (P.setoid A)) (P.setoid (List A)) +Pointwise-≡↔≡ : Inverse (setoid (≡.setoid A)) (≡.setoid (List A)) Pointwise-≡↔≡ = record { to = id ; from = id diff --git a/src/Data/List/Relation/Binary/Pointwise/Properties.agda b/src/Data/List/Relation/Binary/Pointwise/Properties.agda index 7dfba09296..1761dd3cda 100644 --- a/src/Data/List/Relation/Binary/Pointwise/Properties.agda +++ b/src/Data/List/Relation/Binary/Pointwise/Properties.agda @@ -13,7 +13,7 @@ open import Data.List.Base using (List; []; _∷_) open import Level open import Relation.Binary.Core using (REL; _⇒_) open import Relation.Binary.Definitions -import Relation.Binary.PropositionalEquality.Core as P +import Relation.Binary.PropositionalEquality.Core as ≡ open import Relation.Nullary using (yes; no; _×-dec_) import Relation.Nullary.Decidable as Dec @@ -72,6 +72,6 @@ decidable R? (x ∷ xs) (y ∷ ys) = Dec.map′ (uncurry _∷_) uncons (R? x y ×-dec decidable R? xs ys) irrelevant : Irrelevant R → Irrelevant (Pointwise R) -irrelevant irr [] [] = P.refl +irrelevant irr [] [] = ≡.refl irrelevant irr (r ∷ rs) (r₁ ∷ rs₁) = - P.cong₂ _∷_ (irr r r₁) (irrelevant irr rs rs₁) + ≡.cong₂ _∷_ (irr r r₁) (irrelevant irr rs rs₁) diff --git a/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda index 776d065423..8186fc9879 100644 --- a/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda @@ -28,7 +28,7 @@ open import Relation.Unary as U using (Pred) open import Relation.Binary.Core using (Rel; REL; _⇒_) open import Relation.Binary.Definitions using (Trans; Antisym; Irrelevant; Decidable) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; _≢_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; _≢_) private variable @@ -158,7 +158,7 @@ replicate⁺ (s≤s m≤n) r = r ∷ replicate⁺ m≤n r replicate⁻ : ∀ {m n a b} → m ≢ 0 → Prefix R (replicate m a) (replicate n b) → R a b -replicate⁻ {m = zero} {n} m≢0 r = ⊥-elim (m≢0 P.refl) +replicate⁻ {m = zero} {n} m≢0 r = ⊥-elim (m≢0 ≡.refl) replicate⁻ {m = suc m} {suc n} m≢0 rs = Prefix.head rs ------------------------------------------------------------------------ @@ -208,9 +208,9 @@ module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where irrelevant : Irrelevant R → Irrelevant (Prefix R) - irrelevant R-irr [] [] = P.refl + irrelevant R-irr [] [] = ≡.refl irrelevant R-irr (r ∷ rs) (r′ ∷ rs′) = - P.cong₂ _∷_ (R-irr r r′) (irrelevant R-irr rs rs′) + ≡.cong₂ _∷_ (R-irr r r′) (irrelevant R-irr rs rs′) ------------------------------------------------------------------------ -- Decidability diff --git a/src/Data/List/Relation/Binary/Sublist/Heterogeneous.agda b/src/Data/List/Relation/Binary/Sublist/Heterogeneous.agda index 7aa3f4dd5d..f66a29ee5e 100644 --- a/src/Data/List/Relation/Binary/Sublist/Heterogeneous.agda +++ b/src/Data/List/Relation/Binary/Sublist/Heterogeneous.agda @@ -13,7 +13,7 @@ open import Data.List.Relation.Unary.Any using (Any; here; there) open import Level using (_⊔_) open import Relation.Binary.Core using (REL; _⇒_) open import Relation.Binary.Definitions using (_⟶_Respects_; Min) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Unary using (Pred) module Data.List.Relation.Binary.Sublist.Heterogeneous diff --git a/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Properties.agda index 3184ca4420..b760e0287f 100644 --- a/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Sublist/Heterogeneous/Properties.agda @@ -39,7 +39,7 @@ open import Relation.Binary.Definitions using (Reflexive; Trans; Antisym; Decidable; Irrelevant; Irreflexive) open import Relation.Binary.Structures using (IsPreorder; IsPartialOrder; IsDecPartialOrder) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) ------------------------------------------------------------------------ -- Injectivity of constructors @@ -48,15 +48,15 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where ∷-injectiveˡ : ∀ {x y xs ys} {px qx : R x y} {pxs qxs : Sublist R xs ys} → (Sublist R (x ∷ xs) (y ∷ ys) ∋ px ∷ pxs) ≡ (qx ∷ qxs) → px ≡ qx - ∷-injectiveˡ P.refl = P.refl + ∷-injectiveˡ ≡.refl = ≡.refl ∷-injectiveʳ : ∀ {x y xs ys} {px qx : R x y} {pxs qxs : Sublist R xs ys} → (Sublist R (x ∷ xs) (y ∷ ys) ∋ px ∷ pxs) ≡ (qx ∷ qxs) → pxs ≡ qxs - ∷-injectiveʳ P.refl = P.refl + ∷-injectiveʳ ≡.refl = ≡.refl ∷ʳ-injective : ∀ {y xs ys} {pxs qxs : Sublist R xs ys} → (Sublist R xs (y ∷ ys) ∋ y ∷ʳ pxs) ≡ (y ∷ʳ qxs) → pxs ≡ qxs - ∷ʳ-injective P.refl = P.refl + ∷ʳ-injective ≡.refl = ≡.refl module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where @@ -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) (List.reverse-involutive as) (List.reverse-involutive bs) + cast = ≡.subst₂ (Sublist R) (List.reverse-involutive as) (List.reverse-involutive bs) ------------------------------------------------------------------------ -- Inversion lemmas @@ -327,27 +327,27 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} {a as b bs} where module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where Sublist-[]-irrelevant : U.Irrelevant (Sublist R []) - Sublist-[]-irrelevant [] [] = P.refl - Sublist-[]-irrelevant (y ∷ʳ p) (.y ∷ʳ q) = P.cong (y ∷ʳ_) (Sublist-[]-irrelevant p q) + Sublist-[]-irrelevant [] [] = ≡.refl + Sublist-[]-irrelevant (y ∷ʳ p) (.y ∷ʳ q) = ≡.cong (y ∷ʳ_) (Sublist-[]-irrelevant p q) ------------------------------------------------------------------------ -- (to/from)Any is a bijection toAny-injective : ∀ {xs x} {p q : Sublist R [ x ] xs} → toAny p ≡ toAny q → p ≡ q toAny-injective {p = y ∷ʳ p} {y ∷ʳ q} = - P.cong (y ∷ʳ_) ∘′ toAny-injective ∘′ there-injective + ≡.cong (y ∷ʳ_) ∘′ toAny-injective ∘′ there-injective toAny-injective {p = _ ∷ p} {_ ∷ q} = - P.cong₂ (flip _∷_) (Sublist-[]-irrelevant p q) ∘′ here-injective + ≡.cong₂ (flip _∷_) (Sublist-[]-irrelevant p q) ∘′ here-injective fromAny-injective : ∀ {xs x} {p q : Any (R x) xs} → fromAny {R = R} p ≡ fromAny q → p ≡ q - fromAny-injective {p = here px} {here qx} = P.cong here ∘′ ∷-injectiveˡ + fromAny-injective {p = here px} {here qx} = ≡.cong here ∘′ ∷-injectiveˡ fromAny-injective {p = there p} {there q} = - P.cong there ∘′ fromAny-injective ∘′ ∷ʳ-injective + ≡.cong there ∘′ fromAny-injective ∘′ ∷ʳ-injective toAny∘fromAny≗id : ∀ {xs x} (p : Any (R x) xs) → toAny (fromAny {R = R} p) ≡ p - toAny∘fromAny≗id (here px) = P.refl - toAny∘fromAny≗id (there p) = P.cong there (toAny∘fromAny≗id p) + toAny∘fromAny≗id (here px) = ≡.refl + toAny∘fromAny≗id (there p) = ≡.cong there (toAny∘fromAny≗id p) Sublist-[x]-bijection : ∀ {x xs} → (Sublist R [ x ] xs) ⤖ (Any (R x) xs) Sublist-[x]-bijection = mk⤖ (toAny-injective , strictlySurjective⇒surjective < fromAny , toAny∘fromAny≗id >) @@ -360,10 +360,10 @@ module Reflexivity (R-refl : Reflexive R) where reflexive : _≡_ ⇒ Sublist R - reflexive P.refl = fromPointwise (Pw.refl R-refl) + reflexive ≡.refl = fromPointwise (Pw.refl R-refl) refl : Reflexive (Sublist R) - refl = reflexive P.refl + refl = reflexive ≡.refl open Reflexivity public @@ -392,18 +392,18 @@ module Antisymmetry 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 ≡.refl $ begin length (y ∷ ys₁) ≤⟨ length-mono-≤ ss ⟩ 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 ≡.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 ≡.refl $ begin length (y ∷ ys₁) ≤⟨ length-mono-≤ ss ⟩ length xs ≤⟨ length-mono-≤ rs ⟩ length ys₁ ∎ @@ -506,10 +506,10 @@ module Disjointness {a b r} {A : Set a} {B : Set b} {R : REL A B r} where -- Disjoint is proof-irrelevant Disjoint-irrelevant : ∀{xs ys zs} → Irrelevant (Disjoint {R = R} {xs} {ys} {zs}) - Disjoint-irrelevant [] [] = P.refl - Disjoint-irrelevant (y ∷ₙ d₁) (.y ∷ₙ d₂) = P.cong (y ∷ₙ_) (Disjoint-irrelevant d₁ d₂) - Disjoint-irrelevant (x≈y ∷ₗ d₁) (.x≈y ∷ₗ d₂) = P.cong (x≈y ∷ₗ_) (Disjoint-irrelevant d₁ d₂) - Disjoint-irrelevant (x≈y ∷ᵣ d₁) (.x≈y ∷ᵣ d₂) = P.cong (x≈y ∷ᵣ_) (Disjoint-irrelevant d₁ d₂) + Disjoint-irrelevant [] [] = ≡.refl + Disjoint-irrelevant (y ∷ₙ d₁) (.y ∷ₙ d₂) = ≡.cong (y ∷ₙ_) (Disjoint-irrelevant d₁ d₂) + Disjoint-irrelevant (x≈y ∷ₗ d₁) (.x≈y ∷ₗ d₂) = ≡.cong (x≈y ∷ₗ_) (Disjoint-irrelevant d₁ d₂) + Disjoint-irrelevant (x≈y ∷ᵣ d₁) (.x≈y ∷ᵣ d₂) = ≡.cong (x≈y ∷ᵣ_) (Disjoint-irrelevant d₁ d₂) -- Note: DisjointUnion is not proof-irrelevant unless the underlying relation R is. -- The proof is not entirely trivial, thus, we leave it for future work: @@ -525,7 +525,7 @@ module Disjointness {a b r} {A : Set a} {B : Set b} {R : REL A B r} where Disjoint-irrefl′ (y ∷ₙ d) = Disjoint-irrefl′ d Disjoint-irrefl : ∀{x xs ys} → Irreflexive {A = x ∷ xs ⊆ ys } _≡_ Disjoint - Disjoint-irrefl P.refl x with Disjoint-irrefl′ x + Disjoint-irrefl ≡.refl x with Disjoint-irrefl′ x ... | () ∷ _ -- Symmetry diff --git a/src/Data/List/Relation/Binary/Sublist/Setoid.agda b/src/Data/List/Relation/Binary/Sublist/Setoid.agda index 86e16ab80c..83a16444d0 100644 --- a/src/Data/List/Relation/Binary/Sublist/Setoid.agda +++ b/src/Data/List/Relation/Binary/Sublist/Setoid.agda @@ -29,7 +29,7 @@ open import Data.Product.Base using (∃; ∃₂; _×_; _,_; proj₂) open import Relation.Binary.Core using (_⇒_) open import Relation.Binary.Bundles using (Preorder; Poset) open import Relation.Binary.Structures using (IsPreorder; IsPartialOrder) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary using (¬_; Dec; yes; no) open Setoid S renaming (Carrier to A) diff --git a/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda index 17a0c25c77..2090948734 100644 --- a/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda @@ -27,7 +27,7 @@ open import Relation.Nullary.Negation using (contradiction) open import Relation.Binary.Core using (REL; Rel; _⇒_) open import Relation.Binary.Definitions as B using (Trans; Antisym; Irrelevant) -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; _≢_; refl; sym; subst; subst₂) import Data.List.Properties as List @@ -199,10 +199,10 @@ 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 irrelevant : Irrelevant R → Irrelevant (Suffix R) - irrelevant irr (here rs) (here rs₁) = P.cong here $ Pw.irrelevant irr rs rs₁ + irrelevant irr (here rs) (here rs₁) = ≡.cong here $ Pw.irrelevant irr rs rs₁ irrelevant irr (here rs) (there rsuf) = contradiction (Pointwise-length rs) (S[as][bs]⇒∣as∣≢1+∣bs∣ rsuf) irrelevant irr (there rsuf) (here rs) = contradiction (Pointwise-length rs) (S[as][bs]⇒∣as∣≢1+∣bs∣ rsuf) - irrelevant irr (there rsuf) (there rsuf₁) = P.cong there $ irrelevant irr rsuf rsuf₁ + irrelevant irr (there rsuf) (there rsuf₁) = ≡.cong there $ irrelevant irr rsuf rsuf₁ ------------------------------------------------------------------------ -- Decidability diff --git a/src/Data/List/Relation/Ternary/Interleaving.agda b/src/Data/List/Relation/Ternary/Interleaving.agda index 17055652fb..b8201e1072 100644 --- a/src/Data/List/Relation/Ternary/Interleaving.agda +++ b/src/Data/List/Relation/Ternary/Interleaving.agda @@ -16,7 +16,7 @@ open import Data.Product.Base as Product using (∃; ∃₂; _×_; uncurry; _,_; open import Data.Sum.Base using (_⊎_; inj₁; inj₂) open import Function.Base open import Relation.Binary.Core using (REL; _⇒_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) ------------------------------------------------------------------------ -- Definition @@ -60,9 +60,9 @@ module _ {a b c l r} {A : Set a} {B : Set b} {C : Set c} Interleaving _≡_ _≡_ csl csr cs × Pointwise L l csl × Pointwise R r csr break [] = -, [] , [] , [] break (l ∷ˡ sp) = let (_ , eq , pwl , pwr) = break sp in - -, P.refl ∷ˡ eq , l ∷ pwl , pwr + -, ≡.refl ∷ˡ eq , l ∷ pwl , pwr break (r ∷ʳ sp) = let (_ , eq , pwl , pwr) = break sp in - -, P.refl ∷ʳ eq , pwl , r ∷ pwr + -, ≡.refl ∷ʳ eq , pwl , r ∷ pwr -- map diff --git a/src/Data/List/Relation/Unary/All.agda b/src/Data/List/Relation/Unary/All.agda index d9f358df04..bd8677630b 100644 --- a/src/Data/List/Relation/Unary/All.agda +++ b/src/Data/List/Relation/Unary/All.agda @@ -25,8 +25,8 @@ import Relation.Nullary.Decidable as Dec open import Relation.Unary hiding (_∈_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (_Respects_) -open import Relation.Binary.PropositionalEquality.Core as P -import Relation.Binary.PropositionalEquality.Properties as P +open import Relation.Binary.PropositionalEquality.Core as ≡ +import Relation.Binary.PropositionalEquality.Properties as ≡ private variable @@ -124,7 +124,7 @@ module _(S : Setoid a ℓ) {P : Pred (Setoid.Carrier S) p} where tabulateₛ {x ∷ xs} hyp = hyp (here refl₁) ∷ tabulateₛ (hyp ∘ there) tabulate : (∀ {x} → x ∈ₚ xs → P x) → All P xs -tabulate = tabulateₛ (P.setoid _) +tabulate = tabulateₛ (≡.setoid _) self : ∀ {xs : List A} → All (const A) xs self = tabulate (λ {x} _ → x) @@ -213,9 +213,9 @@ universal u [] = [] universal u (x ∷ xs) = u x ∷ universal u xs irrelevant : Irrelevant P → Irrelevant (All P) -irrelevant irr [] [] = P.refl +irrelevant irr [] [] = ≡.refl irrelevant irr (px₁ ∷ pxs₁) (px₂ ∷ pxs₂) = - P.cong₂ _∷_ (irr px₁ px₂) (irrelevant irr pxs₁ pxs₂) + ≡.cong₂ _∷_ (irr px₁ px₂) (irrelevant irr pxs₁ pxs₂) satisfiable : Satisfiable (All P) satisfiable = [] , [] diff --git a/src/Data/List/Relation/Unary/Any/Properties.agda b/src/Data/List/Relation/Unary/Any/Properties.agda index 949f6d9fdb..30dbc6850b 100644 --- a/src/Data/List/Relation/Unary/Any/Properties.agda +++ b/src/Data/List/Relation/Unary/Any/Properties.agda @@ -41,7 +41,7 @@ open import Function.Related.Propositional as Related using (Kind; Related) open import Level using (Level) open import Relation.Binary.Core using (Rel; REL) open import Relation.Binary.Definitions as B -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) @@ -100,13 +100,13 @@ Any-cong {P = P} {Q = Q} {xs = xs} {ys} P↔Q xs≈ys = map-id : ∀ (f : P ⋐ P) → (∀ {x} (p : P x) → f p ≡ p) → (p : Any P xs) → Any.map f p ≡ p -map-id f hyp (here p) = P.cong here (hyp p) -map-id f hyp (there p) = P.cong there $ map-id f hyp p +map-id f hyp (here p) = ≡.cong here (hyp p) +map-id f hyp (there p) = ≡.cong there $ map-id f hyp p map-∘ : ∀ (f : Q ⋐ R) (g : P ⋐ Q) (p : Any P xs) → Any.map (f ∘ g) p ≡ Any.map f (Any.map g p) map-∘ f g (here p) = refl -map-∘ f g (there p) = P.cong there $ map-∘ f g p +map-∘ f g (there p) = ≡.cong there $ map-∘ f g p ------------------------------------------------------------------------ -- Any.lookup @@ -133,16 +133,16 @@ swap-there : ∀ {P : A → B → Set ℓ} → (any : Any (λ x → Any (P x) ys) xs) → swap (Any.map (there {x = x}) any) ≡ there (swap any) swap-there (here pys) = refl -swap-there (there pxys) = P.cong (Any.map there) (swap-there pxys) +swap-there (there pxys) = ≡.cong (Any.map there) (swap-there pxys) swap-invol : ∀ {P : A → B → Set ℓ} → (any : Any (λ x → Any (P x) ys) xs) → swap (swap any) ≡ any swap-invol (here (here px)) = refl swap-invol (here (there pys)) = - P.cong (Any.map there) (swap-invol (here pys)) + ≡.cong (Any.map there) (swap-invol (here pys)) swap-invol (there pxys) = - P.trans (swap-there (swap pxys)) (P.cong there (swap-invol pxys)) + ≡.trans (swap-there (swap pxys)) (≡.cong there (swap-invol pxys)) swap↔ : ∀ {P : A → B → Set ℓ} → Any (λ x → Any (P x) ys) xs ↔ Any (λ y → Any (flip P y) xs) ys @@ -234,7 +234,7 @@ Any-×⁻ pq with Product.map₂ (Product.map₂ find) (find pq) (y , y∈ys , p , q) = find pq in lose x∈xs p , lose y∈ys q) - ≡⟨ P.cong (λ • → let (x , x∈xs , pq) = • + ≡⟨ ≡.cong (λ • → let (x , x∈xs , pq) = • (y , y∈ys , p , q) = find pq in lose x∈xs p , lose y∈ys q) (find∘map p (λ p → Any.map (p ,_) q)) ⟩ @@ -243,7 +243,7 @@ Any-×⁻ pq with Product.map₂ (Product.map₂ find) (find pq) (y , y∈ys , p , q) = find (Any.map (p ,_) q) in lose x∈xs p , lose y∈ys q) - ≡⟨ P.cong (λ • → let (x , x∈xs , p) = find p + ≡⟨ ≡.cong (λ • → let (x , x∈xs , p) = find p (y , y∈ys , p , q) = • in lose x∈xs p , lose y∈ys q) (find∘map q (proj₂ (proj₂ (find p)) ,_)) ⟩ @@ -252,7 +252,7 @@ Any-×⁻ pq with Product.map₂ (Product.map₂ find) (find pq) (y , y∈ys , q) = find q in lose x∈xs p , lose y∈ys q) - ≡⟨ P.cong₂ _,_ (lose∘find p) (lose∘find q) ⟩ + ≡⟨ ≡.cong₂ _,_ (lose∘find p) (lose∘find q) ⟩ (p , q) ∎ @@ -263,15 +263,15 @@ Any-×⁻ pq with Product.map₂ (Product.map₂ find) (find pq) with find pq′ | (λ (f : (proj₁ (find pq′) ≡_) ⋐ _) → map∘find pq′ {f}) ... | (y , y∈ys , p , q) | lem₂ - rewrite P.sym $ map-∘ {R = λ x → Any (λ y → P x × Q y) ys} + rewrite ≡.sym $ map-∘ {R = λ x → Any (λ y → P x × Q y) ys} (λ p → Any.map (λ q → p , q) (lose y∈ys q)) - (λ y → P.subst P y p) + (λ y → ≡.subst P y p) x∈xs = lem₁ _ helper where helper : Any.map (λ q → p , q) (lose y∈ys q) ≡ pq′ - helper rewrite P.sym $ map-∘ (λ q → p , q) - (λ y → P.subst Q y q) + helper rewrite ≡.sym $ map-∘ (λ q → p , q) + (λ y → ≡.subst Q y q) y∈ys = lem₂ _ refl @@ -317,12 +317,12 @@ module _ {f : A → B} where map⁺∘map⁻ : (p : Any P (List.map f xs)) → map⁺ (map⁻ p) ≡ p map⁺∘map⁻ {xs = x ∷ xs} (here p) = refl - map⁺∘map⁻ {xs = x ∷ xs} (there p) = P.cong there (map⁺∘map⁻ p) + map⁺∘map⁻ {xs = x ∷ xs} (there p) = ≡.cong there (map⁺∘map⁻ p) map⁻∘map⁺ : ∀ (P : Pred B p) → (p : Any (P ∘ f) xs) → map⁻ {P = P} (map⁺ p) ≡ p map⁻∘map⁺ P (here p) = refl - map⁻∘map⁺ P (there p) = P.cong there (map⁻∘map⁺ P p) + map⁻∘map⁺ P (there p) = ≡.cong there (map⁻∘map⁺ P p) map↔ : Any (P ∘ f) xs ↔ Any P (List.map f xs) map↔ = mk↔ₛ′ map⁺ map⁻ map⁺∘map⁻ (map⁻∘map⁺ _) @@ -363,8 +363,8 @@ module _ {P : A → Set p} where ++⁺∘++⁻ [] p = refl ++⁺∘++⁻ (x ∷ xs) (here p) = refl ++⁺∘++⁻ (x ∷ xs) (there p) with ++⁻ xs p | ++⁺∘++⁻ xs p - ... | inj₁ p′ | ih = P.cong there ih - ... | inj₂ p′ | ih = P.cong there ih + ... | inj₁ p′ | ih = ≡.cong there ih + ... | inj₂ p′ | ih = ≡.cong there ih ++⁻∘++⁺ : ∀ xs {ys} (p : Any P xs ⊎ Any P ys) → ++⁻ xs ([ ++⁺ˡ , ++⁺ʳ xs ]′ p) ≡ p @@ -441,7 +441,7 @@ module _ {P : A → Set p} where concat⁻∘concat⁺ (here p) = concat⁻∘++⁺ˡ _ p concat⁻∘concat⁺ (there {x = xs} {xs = xss} p) rewrite concat⁻∘++⁺ʳ xs xss (concat⁺ p) = - P.cong there $ concat⁻∘concat⁺ p + ≡.cong there $ concat⁻∘concat⁺ p concat↔ : ∀ {xss} → Any (Any P) xss ↔ Any P (concat xss) concat↔ {xss} = mk↔ₛ′ concat⁺ (concat⁻ xss) (concat⁺∘concat⁻ xss) concat⁻∘concat⁺ @@ -496,7 +496,7 @@ applyUpTo⁻ f {suc n} (there p) with applyUpTo⁻ (f ∘ suc) p applyDownFrom⁺ : ∀ f {i n} → P (f i) → i < n → Any P (applyDownFrom f n) applyDownFrom⁺ f {i} {suc n} p (s≤s i≤n) with i ≟ n -... | yes P.refl = here p +... | yes ≡.refl = here p ... | no i≢n = there (applyDownFrom⁺ f p (≤∧≢⇒< i≤n i≢n)) applyDownFrom⁻ : ∀ f {n} → Any P (applyDownFrom f n) → @@ -611,7 +611,7 @@ module _ {P : B → Set p} where mapWith∈⁺ f (mapWith∈⁻ xs f p) ≡ p to∘from (y ∷ xs) f (here p) = refl to∘from (y ∷ xs) f (there p) = - P.cong there $ to∘from xs (f ∘ there) p + ≡.cong there $ to∘from xs (f ∘ there) p ------------------------------------------------------------------------ -- reverse @@ -686,7 +686,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where Any (λ f → Any (P ∘ f) xs) fs ↔⟨ Any-cong (λ _ → Any-cong (λ _ → pure↔) (_ ∎)) (_ ∎) ⟩ Any (λ f → Any (Any P ∘ pure ∘ f) xs) fs ↔⟨ Any-cong (λ _ → >>=↔ ) (_ ∎) ⟩ Any (λ f → Any P (xs >>= pure ∘ f)) fs ↔⟨ >>=↔ ⟩ - Any P (fs >>= λ f → xs >>= λ x → pure (f x)) ≡⟨ P.cong (Any P) (Listₑ.Applicative.unfold-⊛ fs xs) ⟨ + Any P (fs >>= λ f → xs >>= λ x → pure (f x)) ≡⟨ ≡.cong (Any P) (Listₑ.Applicative.unfold-⊛ fs xs) ⟨ Any P (fs ⊛ xs) ∎ where open Related.EquationalReasoning @@ -706,7 +706,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where Any (λ x → Any (λ y → P (x , y)) ys) xs ↔⟨ pure↔ ⟩ Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (pure _,_) ↔⟨ ⊛↔ ⟩ Any (λ x, → Any (P ∘ x,) ys) (pure _,_ ⊛ xs) ↔⟨ ⊛↔ ⟩ - Any P (pure _,_ ⊛ xs ⊛ ys) ≡⟨ P.cong (Any P ∘′ (_⊛ ys)) (Listₑ.Applicative.unfold-<$> _,_ xs) ⟨ + Any P (pure _,_ ⊛ xs ⊛ ys) ≡⟨ ≡.cong (Any P ∘′ (_⊛ ys)) (Listₑ.Applicative.unfold-<$> _,_ xs) ⟨ Any P (xs ⊗ ys) ∎ where open Related.EquationalReasoning diff --git a/src/Data/List/Relation/Unary/Enumerates/Setoid/Properties.agda b/src/Data/List/Relation/Unary/Enumerates/Setoid/Properties.agda index f5e0a53c5c..f1a95b1e14 100644 --- a/src/Data/List/Relation/Unary/Enumerates/Setoid/Properties.agda +++ b/src/Data/List/Relation/Unary/Enumerates/Setoid/Properties.agda @@ -23,7 +23,7 @@ open import Function.Definitions using (Surjective) open import Function.Consequences using (strictlySurjective⇒surjective) open import Level open import Relation.Binary.Bundles using (Setoid; DecSetoid) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Properties.Setoid using (respʳ-flip) module Data.List.Relation.Unary.Enumerates.Setoid.Properties where @@ -89,4 +89,4 @@ module _ (S : Setoid a ℓ₁) where lookup-surjective : ∀ {xs} → IsEnumeration S xs → Surjective _≡_ _≈_ (lookup xs) lookup-surjective _∈xs = strictlySurjective⇒surjective - trans (λ { P.refl → refl}) (λ y → index (y ∈xs) , sym (lookup-index (y ∈xs))) + trans (λ { ≡.refl → refl}) (λ y → index (y ∈xs) , sym (lookup-index (y ∈xs))) diff --git a/src/Data/Maybe/Relation/Binary/Connected.agda b/src/Data/Maybe/Relation/Binary/Connected.agda index 8c174c9ea1..7bf8f33310 100644 --- a/src/Data/Maybe/Relation/Binary/Connected.agda +++ b/src/Data/Maybe/Relation/Binary/Connected.agda @@ -13,7 +13,7 @@ open import Data.Maybe.Base using (Maybe; just; nothing) open import Function.Bundles using (_⇔_; mk⇔) open import Relation.Binary.Core using (REL; _⇒_) open import Relation.Binary.Definitions using (Reflexive; Sym; Decidable) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary import Relation.Nullary.Decidable as Dec @@ -52,7 +52,7 @@ refl R-refl {just _} = just R-refl refl R-refl {nothing} = nothing reflexive : _≡_ ⇒ R → _≡_ ⇒ Connected R -reflexive reflexive P.refl = refl (reflexive P.refl) +reflexive reflexive ≡.refl = refl (reflexive ≡.refl) sym : Sym R S → Sym (Connected R) (Connected S) sym R-sym (just p) = just (R-sym p) diff --git a/src/Data/Maybe/Relation/Binary/Pointwise.agda b/src/Data/Maybe/Relation/Binary/Pointwise.agda index 3da2abd9f1..073ac74a81 100644 --- a/src/Data/Maybe/Relation/Binary/Pointwise.agda +++ b/src/Data/Maybe/Relation/Binary/Pointwise.agda @@ -17,7 +17,7 @@ open import Relation.Binary.Core using (REL; Rel; _⇒_) open import Relation.Binary.Bundles using (Setoid; DecSetoid) open import Relation.Binary.Definitions using (Reflexive; Sym; Trans; Decidable) open import Relation.Binary.Structures using (IsEquivalence; IsDecEquivalence) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary open import Relation.Unary using (_⊆_) import Relation.Nullary.Decidable as Dec @@ -43,10 +43,10 @@ module _ {a b ℓ} {A : Set a} {B : Set b} {R : REL A B ℓ} where just-equivalence = mk⇔ just drop-just nothing-inv : ∀ {x} → Pointwise R nothing x → x ≡ nothing - nothing-inv nothing = P.refl + nothing-inv nothing = ≡.refl just-inv : ∀ {x y} → Pointwise R (just x) y → ∃ λ z → y ≡ just z × R x z - just-inv (just r) = -, P.refl , r + just-inv (just r) = -, ≡.refl , r ------------------------------------------------------------------------ -- Relational properties @@ -58,7 +58,7 @@ module _ {a r} {A : Set a} {R : Rel A r} where refl R-refl {nothing} = nothing reflexive : _≡_ ⇒ R → _≡_ ⇒ Pointwise R - reflexive reflexive P.refl = refl (reflexive P.refl) + reflexive reflexive ≡.refl = refl (reflexive ≡.refl) module _ {a b r₁ r₂} {A : Set a} {B : Set b} {R : REL A B r₁} {S : REL B A r₂} where diff --git a/src/Data/Maybe/Relation/Unary/All.agda b/src/Data/Maybe/Relation/Unary/All.agda index 9a57aa62f5..31e4d8fd83 100644 --- a/src/Data/Maybe/Relation/Unary/All.agda +++ b/src/Data/Maybe/Relation/Unary/All.agda @@ -16,7 +16,7 @@ open import Data.Product.Base as Product using (_,_) open import Function.Base using (id; _∘′_) open import Function.Bundles using (_⇔_; mk⇔) open import Level -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; cong) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; cong) open import Relation.Unary open import Relation.Nullary hiding (Irrelevant) import Relation.Nullary.Decidable as Dec @@ -114,7 +114,7 @@ module _ {a p} {A : Set a} {P : Pred A p} where irrelevant : Irrelevant P → Irrelevant (All P) irrelevant P-irrelevant (just p) (just q) = cong just (P-irrelevant p q) - irrelevant P-irrelevant nothing nothing = P.refl + irrelevant P-irrelevant nothing nothing = ≡.refl satisfiable : Satisfiable (All P) satisfiable = nothing , nothing diff --git a/src/Data/Maybe/Relation/Unary/Any.agda b/src/Data/Maybe/Relation/Unary/Any.agda index 767a5624cc..d8cf44376f 100644 --- a/src/Data/Maybe/Relation/Unary/Any.agda +++ b/src/Data/Maybe/Relation/Unary/Any.agda @@ -13,7 +13,7 @@ open import Data.Product.Base as Product using (∃; _,_; -,_) open import Function.Base using (id) open import Function.Bundles using (_⇔_; mk⇔) open import Level -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; cong) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; cong) open import Relation.Unary open import Relation.Nullary hiding (Irrelevant) import Relation.Nullary.Decidable as Dec diff --git a/src/Data/Nat/Coprimality.agda b/src/Data/Nat/Coprimality.agda index f513e82fa4..16a9cf9897 100644 --- a/src/Data/Nat/Coprimality.agda +++ b/src/Data/Nat/Coprimality.agda @@ -19,7 +19,7 @@ open import Data.Product.Base as Prod open import Data.Sum.Base as Sum using (inj₁; inj₂) open import Function.Base using (_∘_) open import Level using (0ℓ) -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; _≢_; refl; trans; cong; subst) open import Relation.Nullary as Nullary using (¬_; contradiction; map′) open import Relation.Binary.Core using (Rel) @@ -47,7 +47,7 @@ coprime⇒GCD≡1 {m} {n} coprime = GCD.is (1∣ m , 1∣ n) (∣-reflexive ∘ GCD≡1⇒coprime : GCD m n 1 → Coprime m n GCD≡1⇒coprime g cd with divides q eq ← GCD.greatest g cd - = m*n≡1⇒n≡1 q _ (P.sym eq) + = m*n≡1⇒n≡1 q _ (≡.sym eq) coprime⇒gcd≡1 : Coprime m n → gcd m n ≡ 1 coprime⇒gcd≡1 coprime = GCD.unique (gcd-GCD _ _) (coprime⇒GCD≡1 coprime) diff --git a/src/Data/Nat/GCD.agda b/src/Data/Nat/GCD.agda index b88947c707..26d44549fb 100644 --- a/src/Data/Nat/GCD.agda +++ b/src/Data/Nat/GCD.agda @@ -22,7 +22,7 @@ open import Function.Base using (_$_; _∘_) open import Induction using (build) open import Induction.Lexicographic using (_⊗_; [_⊗_]) open import Relation.Binary.Definitions using (tri<; tri>; tri≈; Symmetric) -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; _≢_; subst; cong) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) @@ -79,7 +79,7 @@ gcd[m,n]∣m m n with <-cmp m n gcd[m,n]∣n : ∀ m n → gcd m n ∣ n gcd[m,n]∣n m n with <-cmp m n ... | tri< n _ _ m; tri≈) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) ------------------------------------------------------------------------ -- Definition diff --git a/src/Data/Sum/Relation/Binary/Pointwise.agda b/src/Data/Sum/Relation/Binary/Pointwise.agda index 943b6f5e49..6d6c84a253 100644 --- a/src/Data/Sum/Relation/Binary/Pointwise.agda +++ b/src/Data/Sum/Relation/Binary/Pointwise.agda @@ -17,8 +17,8 @@ open import Function.Bundles using (Inverse; mk↔) open import Relation.Nullary import Relation.Nullary.Decidable as Dec open import Relation.Binary -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 @@ -201,14 +201,14 @@ _⊎ₛ_ = ⊎-setoid -- decomposed using Pointwise Pointwise-≡⇒≡ : (Pointwise _≡_ _≡_) ⇒ _≡_ {A = A ⊎ B} -Pointwise-≡⇒≡ (inj₁ x) = P.cong inj₁ x -Pointwise-≡⇒≡ (inj₂ x) = P.cong inj₂ x +Pointwise-≡⇒≡ (inj₁ x) = ≡.cong inj₁ x +Pointwise-≡⇒≡ (inj₂ x) = ≡.cong inj₂ x ≡⇒Pointwise-≡ : _≡_ {A = A ⊎ B} ⇒ (Pointwise _≡_ _≡_) -≡⇒Pointwise-≡ P.refl = ⊎-refl P.refl P.refl +≡⇒Pointwise-≡ ≡.refl = ⊎-refl ≡.refl ≡.refl Pointwise-≡↔≡ : (A : Set a) (B : Set b) → - Inverse (P.setoid A ⊎ₛ P.setoid B) (P.setoid (A ⊎ B)) + Inverse (≡.setoid A ⊎ₛ ≡.setoid B) (≡.setoid (A ⊎ B)) Pointwise-≡↔≡ _ _ = record { to = id ; from = id diff --git a/src/Data/Vec/Functional/Relation/Binary/Equality/Setoid.agda b/src/Data/Vec/Functional/Relation/Binary/Equality/Setoid.agda index 5853bd1f68..b14a2eaee2 100644 --- a/src/Data/Vec/Functional/Relation/Binary/Equality/Setoid.agda +++ b/src/Data/Vec/Functional/Relation/Binary/Equality/Setoid.agda @@ -17,7 +17,7 @@ open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) module Data.Vec.Functional.Relation.Binary.Equality.Setoid {a ℓ} (S : Setoid a ℓ) where @@ -41,7 +41,7 @@ _≋_ = Pointwise _≈_ ≋-refl {n} = PW.refl {R = _≈_} refl ≋-reflexive : ∀ {n} → _≡_ ⇒ (_≋_ {n = n}) -≋-reflexive P.refl = ≋-refl +≋-reflexive ≡.refl = ≋-refl ≋-sym : ∀ {n} → Symmetric (_≋_ {n = n}) ≋-sym = PW.sym {R = _≈_} sym diff --git a/src/Data/Vec/Properties/WithK.agda b/src/Data/Vec/Properties/WithK.agda index 7ca8cc2c91..bcd914ad51 100644 --- a/src/Data/Vec/Properties/WithK.agda +++ b/src/Data/Vec/Properties/WithK.agda @@ -12,8 +12,8 @@ module Data.Vec.Properties.WithK where open import Data.Nat.Base open import Data.Nat.Properties using (+-assoc) open import Data.Vec.Base -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; refl) -open import Relation.Binary.HeterogeneousEquality as H using (_≅_; refl) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +open import Relation.Binary.HeterogeneousEquality as ≅ using (_≅_; refl) ------------------------------------------------------------------------ -- _[_]=_ @@ -24,7 +24,7 @@ module _ {a} {A : Set a} where (p q : xs [ i ]= x) → p ≡ q []=-irrelevant here here = refl []=-irrelevant (there xs[i]=x) (there xs[i]=x′) = - P.cong there ([]=-irrelevant xs[i]=x xs[i]=x′) + ≡.cong there ([]=-irrelevant xs[i]=x xs[i]=x′) ------------------------------------------------------------------------ -- _++_ @@ -35,7 +35,7 @@ module _ {a} {A : Set a} where (xs ++ ys) ++ zs ≅ xs ++ (ys ++ zs) ++-assoc [] ys zs = refl ++-assoc {suc m} (x ∷ xs) ys zs = - H.icong (Vec A) (+-assoc m _ _) (x ∷_) (++-assoc xs ys zs) + ≅.icong (Vec A) (+-assoc m _ _) (x ∷_) (++-assoc xs ys zs) ------------------------------------------------------------------------ -- foldr diff --git a/src/Data/Vec/Recursive/Properties.agda b/src/Data/Vec/Recursive/Properties.agda index d036134a45..df34cab770 100644 --- a/src/Data/Vec/Recursive/Properties.agda +++ b/src/Data/Vec/Recursive/Properties.agda @@ -14,7 +14,7 @@ open import Data.Product.Base open import Data.Vec.Recursive open import Data.Vec.Base using (Vec; _∷_) open import Function.Bundles using (_↔_; mk↔ₛ′) -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) open ≡-Reasoning @@ -28,31 +28,31 @@ private -- Basic proofs cons-head-tail-identity : ∀ n (as : A ^ suc n) → cons n (head n as) (tail n as) ≡ as -cons-head-tail-identity 0 as = P.refl -cons-head-tail-identity (suc n) as = P.refl +cons-head-tail-identity 0 as = ≡.refl +cons-head-tail-identity (suc n) as = ≡.refl head-cons-identity : ∀ n a (as : A ^ n) → head n (cons n a as) ≡ a -head-cons-identity 0 a as = P.refl -head-cons-identity (suc n) a as = P.refl +head-cons-identity 0 a as = ≡.refl +head-cons-identity (suc n) a as = ≡.refl tail-cons-identity : ∀ n a (as : A ^ n) → tail n (cons n a as) ≡ as -tail-cons-identity 0 a as = P.refl -tail-cons-identity (suc n) a as = P.refl +tail-cons-identity 0 a as = ≡.refl +tail-cons-identity (suc n) a as = ≡.refl append-cons : ∀ m n a (xs : A ^ m) ys → append (suc m) n (cons m a xs) ys ≡ cons (m + n) a (append m n xs ys) -append-cons 0 n a xs ys = P.refl -append-cons (suc m) n a xs ys = P.refl +append-cons 0 n a xs ys = ≡.refl +append-cons (suc m) n a xs ys = ≡.refl append-splitAt-identity : ∀ m n (as : A ^ (m + n)) → uncurry (append m n) (splitAt m n as) ≡ as -append-splitAt-identity 0 n as = P.refl +append-splitAt-identity 0 n as = ≡.refl append-splitAt-identity (suc m) n as = begin let x = head (m + n) as in let (xs , ys) = splitAt m n (tail (m + n) as) in append (suc m) n (cons m (head (m + n) as) xs) ys ≡⟨ append-cons m n x xs ys ⟩ cons (m + n) x (append m n xs ys) - ≡⟨ P.cong (cons (m + n) x) (append-splitAt-identity m n (tail (m + n) as)) ⟩ + ≡⟨ ≡.cong (cons (m + n) x) (append-splitAt-identity m n (tail (m + n) as)) ⟩ cons (m + n) x (tail (m + n) as) ≡⟨ cons-head-tail-identity (m + n) as ⟩ as @@ -62,21 +62,21 @@ append-splitAt-identity (suc m) n as = begin -- Conversion to and from Vec fromVec∘toVec : ∀ n (xs : A ^ n) → fromVec (toVec n xs) ≡ xs -fromVec∘toVec 0 _ = P.refl +fromVec∘toVec 0 _ = ≡.refl fromVec∘toVec (suc n) xs = begin cons n (head n xs) (fromVec (toVec n (tail n xs))) - ≡⟨ P.cong (cons n (head n xs)) (fromVec∘toVec n (tail n xs)) ⟩ + ≡⟨ ≡.cong (cons n (head n xs)) (fromVec∘toVec n (tail n xs)) ⟩ cons n (head n xs) (tail n xs) ≡⟨ cons-head-tail-identity n xs ⟩ xs ∎ toVec∘fromVec : ∀ {n} (xs : Vec A n) → toVec n (fromVec xs) ≡ xs -toVec∘fromVec Vec.[] = P.refl +toVec∘fromVec Vec.[] = ≡.refl toVec∘fromVec {n = suc n} (x Vec.∷ xs) = begin head n (cons n x (fromVec xs)) Vec.∷ toVec n (tail n (cons n x (fromVec xs))) - ≡⟨ P.cong₂ (λ x xs → x Vec.∷ toVec n xs) hd-prf tl-prf ⟩ + ≡⟨ ≡.cong₂ (λ x xs → x Vec.∷ toVec n xs) hd-prf tl-prf ⟩ x Vec.∷ toVec n (fromVec xs) - ≡⟨ P.cong (x Vec.∷_) (toVec∘fromVec xs) ⟩ + ≡⟨ ≡.cong (x Vec.∷_) (toVec∘fromVec xs) ⟩ x Vec.∷ xs ∎ where diff --git a/src/Data/Vec/Relation/Binary/Lex/Core.agda b/src/Data/Vec/Relation/Binary/Lex/Core.agda index 7b35081911..ed23c7cf6a 100644 --- a/src/Data/Vec/Relation/Binary/Lex/Core.agda +++ b/src/Data/Vec/Relation/Binary/Lex/Core.agda @@ -22,7 +22,7 @@ open import Relation.Binary.Core using (Rel; REL) open import Relation.Binary.Definitions using (Transitive; Symmetric; Asymmetric; Antisymmetric; Irreflexive; Trans; _Respects₂_; _Respectsˡ_; _Respectsʳ_; Decidable; Irrelevant) open import Relation.Binary.Structures using (IsPartialEquivalence) -open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl; cong) import Relation.Nullary as Nullary open import Relation.Nullary.Decidable as Dec using (Dec; yes; no; _×-dec_; _⊎-dec_) @@ -101,9 +101,9 @@ module _ {P : Set} {_≈_ : Rel A ℓ₁} {_≺_ : Rel A ℓ₂} where transitive′ : ∀ {m n o P₂} → Trans (Lex P _≈_ _≺_ {m} {n}) (Lex P₂ _≈_ _≺_ {n} {o}) (Lex (P × P₂) _≈_ _≺_) transitive′ (base p₁) (base p₂) = base (p₁ , p₂) - transitive′ (this x≺y m≡n) (this y≺z n≡o) = this (≺-trans x≺y y≺z) (P.trans m≡n n≡o) - transitive′ (this x≺y m≡n) (next y≈z ys; tri<) open import Relation.Binary.Consequences open import Relation.Binary.Construct.On as On using (wellFounded) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Level using (Level; _⊔_) private @@ -103,8 +103,8 @@ module _ {_≈_ : Rel A ℓ₁} {_≺_ : Rel A ℓ₂} where <-cmp : ∀ {n} → Trichotomous _≋_ (_<_ {n}) <-cmp [] [] = tri≈ ¬[]<[] [] ¬[]<[] <-cmp (x ∷ xs) (y ∷ ys) with ≺-cmp x y - ... | tri< x≺y x≉y x⊁y = tri< (this x≺y P.refl) (x≉y ∘ head) (≰-this (x≉y ∘ ≈-sym) x⊁y) - ... | tri> x⊀y x≉y x≻y = tri> (≰-this x≉y x⊀y) (x≉y ∘ head) (this x≻y P.refl) + ... | tri< x≺y x≉y x⊁y = tri< (this x≺y ≡.refl) (x≉y ∘ head) (≰-this (x≉y ∘ ≈-sym) x⊁y) + ... | tri> x⊀y x≉y x≻y = tri> (≰-this x≉y x⊀y) (x≉y ∘ head) (this x≻y ≡.refl) ... | tri≈ x⊀y x≈y x⊁y with <-cmp xs ys ... | tri< xs _ _ x≻y = inj₂ (this x≻y P.refl) + ... | tri< x≺y _ _ = inj₁ (this x≺y ≡.refl) + ... | tri> _ _ x≻y = inj₂ (this x≻y ≡.refl) ... | tri≈ _ x≈y _ with ≤-total xs ys ... | inj₁ xsys = inj₂ (next (≈-sym x≈y) xs>ys) diff --git a/src/Data/Vec/Relation/Binary/Pointwise/Extensional.agda b/src/Data/Vec/Relation/Binary/Pointwise/Extensional.agda index 5dcaa6b900..928bae5176 100644 --- a/src/Data/Vec/Relation/Binary/Pointwise/Extensional.agda +++ b/src/Data/Vec/Relation/Binary/Pointwise/Extensional.agda @@ -22,7 +22,7 @@ open import Relation.Binary.Core using (Rel; REL; _⇒_; _=[_]⇒_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence; IsDecEquivalence) open import Relation.Binary.Definitions using (Reflexive; Sym; Trans; Decidable) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Construct.Closure.Transitive as Plus hiding (equivalent; map) open import Relation.Nullary @@ -134,12 +134,12 @@ isDecEquivalence decEquiv = record -- Pointwise _≡_ is equivalent to _≡_. Pointwise-≡⇒≡ : ∀ {n} {xs ys : Vec A n} → Pointwise _≡_ xs ys → xs ≡ ys -Pointwise-≡⇒≡ {xs = []} {[]} (ext app) = P.refl +Pointwise-≡⇒≡ {xs = []} {[]} (ext app) = ≡.refl Pointwise-≡⇒≡ {xs = x ∷ xs} {y ∷ ys} xs∼ys = - P.cong₂ _∷_ (head xs∼ys) (Pointwise-≡⇒≡ (tail xs∼ys)) + ≡.cong₂ _∷_ (head xs∼ys) (Pointwise-≡⇒≡ (tail xs∼ys)) ≡⇒Pointwise-≡ : ∀ {n} {xs ys : Vec A n} → xs ≡ ys → Pointwise _≡_ xs ys -≡⇒Pointwise-≡ P.refl = refl P.refl +≡⇒Pointwise-≡ ≡.refl = refl ≡.refl Pointwise-≡↔≡ : ∀ {n} {xs ys : Vec A n} → Pointwise _≡_ xs ys ⇔ xs ≡ ys Pointwise-≡↔≡ {ℓ} {A} = mk⇔ Pointwise-≡⇒≡ ≡⇒Pointwise-≡ diff --git a/src/Data/Vec/Relation/Binary/Pointwise/Inductive.agda b/src/Data/Vec/Relation/Binary/Pointwise/Inductive.agda index c6835b73e1..17cc144319 100644 --- a/src/Data/Vec/Relation/Binary/Pointwise/Inductive.agda +++ b/src/Data/Vec/Relation/Binary/Pointwise/Inductive.agda @@ -22,7 +22,7 @@ open import Relation.Binary.Structures using (IsEquivalence; IsDecEquivalence) open import Relation.Binary.Definitions using (Trans; Decidable; Reflexive; Sym) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary.Decidable using (yes; no; _×-dec_; map′) open import Relation.Unary using (Pred) @@ -52,8 +52,8 @@ data Pointwise {a b ℓ} {A : Set a} {B : Set b} (_∼_ : REL A B ℓ) : length-equal : ∀ {m n} {_∼_ : REL A B ℓ} {xs : Vec A m} {ys : Vec B n} → Pointwise _∼_ xs ys → m ≡ n -length-equal [] = P.refl -length-equal (_ ∷ xs∼ys) = P.cong suc (length-equal xs∼ys) +length-equal [] = ≡.refl +length-equal (_ ∷ xs∼ys) = ≡.cong suc (length-equal xs∼ys) ------------------------------------------------------------------------ -- Operations @@ -261,11 +261,11 @@ module _ {P : Pred A ℓ} where -- Pointwise _≡_ is equivalent to _≡_ Pointwise-≡⇒≡ : ∀ {n} {xs ys : Vec A n} → Pointwise _≡_ xs ys → xs ≡ ys -Pointwise-≡⇒≡ [] = P.refl -Pointwise-≡⇒≡ (P.refl ∷ xs∼ys) = P.cong (_ ∷_) (Pointwise-≡⇒≡ xs∼ys) +Pointwise-≡⇒≡ [] = ≡.refl +Pointwise-≡⇒≡ (≡.refl ∷ xs∼ys) = ≡.cong (_ ∷_) (Pointwise-≡⇒≡ xs∼ys) ≡⇒Pointwise-≡ : ∀ {n} {xs ys : Vec A n} → xs ≡ ys → Pointwise _≡_ xs ys -≡⇒Pointwise-≡ P.refl = refl P.refl +≡⇒Pointwise-≡ ≡.refl = refl ≡.refl Pointwise-≡↔≡ : ∀ {n} {xs ys : Vec A n} → Pointwise _≡_ xs ys ⇔ xs ≡ ys Pointwise-≡↔≡ = mk⇔ Pointwise-≡⇒≡ ≡⇒Pointwise-≡ diff --git a/src/Data/Vec/Relation/Unary/All.agda b/src/Data/Vec/Relation/Unary/All.agda index 4cc1805c25..8371d36c78 100644 --- a/src/Data/Vec/Relation/Unary/All.agda +++ b/src/Data/Vec/Relation/Unary/All.agda @@ -21,7 +21,7 @@ open import Relation.Nullary.Decidable as Dec using (_×-dec_; yes; no) open import Relation.Unary hiding (_∈_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (_Respects_) -open import Relation.Binary.PropositionalEquality.Core as P using (subst) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (subst) private variable @@ -94,7 +94,7 @@ lookupWith : ∀[ P ⇒ Q ⇒ R ] → All P xs → (i : Any Q xs) → R (Any.loo 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 +lookup pxs = lookupWith (λ { px ≡.refl → px }) pxs module _(S : Setoid a ℓ) {P : Pred (Setoid.Carrier S) p} where open Setoid S renaming (sym to sym₁) @@ -115,9 +115,9 @@ universal u [] = [] universal u (x ∷ xs) = u x ∷ universal u xs irrelevant : Irrelevant P → ∀ {n} → Irrelevant (All P {n}) -irrelevant irr [] [] = P.refl +irrelevant irr [] [] = ≡.refl irrelevant irr (px₁ ∷ pxs₁) (px₂ ∷ pxs₂) = - P.cong₂ _∷_ (irr px₁ px₂) (irrelevant irr pxs₁ pxs₂) + ≡.cong₂ _∷_ (irr px₁ px₂) (irrelevant irr pxs₁ pxs₂) satisfiable : Satisfiable P → ∀ {n} → Satisfiable (All P {n}) satisfiable (x , p) {zero} = [] , [] diff --git a/src/Data/Vec/Relation/Unary/Any/Properties.agda b/src/Data/Vec/Relation/Unary/Any/Properties.agda index d330ea5cef..e95ae39930 100644 --- a/src/Data/Vec/Relation/Unary/Any/Properties.agda +++ b/src/Data/Vec/Relation/Unary/Any/Properties.agda @@ -30,7 +30,7 @@ open import Relation.Nullary.Negation using (¬_) open import Relation.Unary hiding (_∈_) open import Relation.Binary.Core using (Rel) open import Relation.Binary.Definitions using (_Respects_) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) private variable @@ -93,15 +93,15 @@ module _ {P : Pred A p} where map-id : ∀ {P : Pred A p} (f : P ⊆ P) {n xs} → (∀ {x} (p : P x) → f p ≡ p) → (p : Any P {n} xs) → Any.map f p ≡ p -map-id f hyp (here p) = P.cong here (hyp p) -map-id f hyp (there p) = P.cong there $ map-id f hyp p +map-id f hyp (here p) = ≡.cong here (hyp p) +map-id f hyp (there p) = ≡.cong there $ map-id f hyp p map-∘ : ∀ {P : Pred A p} {Q : A → Set q} {R : A → Set r} (f : Q ⊆ R) (g : P ⊆ Q) {n xs} (p : Any P {n} xs) → Any.map (f ∘ g) p ≡ Any.map f (Any.map g p) map-∘ f g (here p) = refl -map-∘ f g (there p) = P.cong there $ map-∘ f g p +map-∘ f g (there p) = ≡.cong there $ map-∘ f g p ------------------------------------------------------------------------ -- Swapping @@ -117,7 +117,7 @@ module _ {P : A → B → Set ℓ} where swap-there : ∀ {n m x xs ys} → (any : Any (λ x → Any (P x) {n} ys) {m} xs) → swap (Any.map (there {x = x}) any) ≡ there (swap any) swap-there (here pys) = refl - swap-there (there pxys) = P.cong (Any.map there) (swap-there pxys) + swap-there (there pxys) = ≡.cong (Any.map there) (swap-there pxys) module _ {P : A → B → Set ℓ} where @@ -125,9 +125,9 @@ module _ {P : A → B → Set ℓ} where (any : Any (λ x → Any (P x) ys) xs) → swap (swap any) ≡ any swap-invol (here (here _)) = refl - swap-invol (here (there pys)) = P.cong (Any.map there) (swap-invol (here pys)) - swap-invol (there pxys) = P.trans (swap-there (swap pxys)) - $ P.cong there (swap-invol pxys) + swap-invol (here (there pys)) = ≡.cong (Any.map there) (swap-invol (here pys)) + swap-invol (there pxys) = ≡.trans (swap-there (swap pxys)) + $ ≡.cong there (swap-invol pxys) module _ {P : A → B → Set ℓ} where @@ -237,12 +237,12 @@ module _ {f : A → B} where map⁺∘map⁻ : ∀ {P : Pred B p} {n} {xs : Vec A n} → (p : Any P (map f xs)) → map⁺ (map⁻ p) ≡ p map⁺∘map⁻ {xs = x ∷ xs} (here p) = refl - map⁺∘map⁻ {xs = x ∷ xs} (there p) = P.cong there (map⁺∘map⁻ p) + map⁺∘map⁻ {xs = x ∷ xs} (there p) = ≡.cong there (map⁺∘map⁻ p) map⁻∘map⁺ : ∀ (P : Pred B p) {n} {xs : Vec A n} → (p : Any (P ∘ f) xs) → map⁻ {P = P} (map⁺ p) ≡ p map⁻∘map⁺ P (here p) = refl - map⁻∘map⁺ P (there p) = P.cong there (map⁻∘map⁺ P p) + map⁻∘map⁺ P (there p) = ≡.cong there (map⁻∘map⁺ P p) map↔ : ∀ {P : Pred B p} {n} {xs : Vec A n} → Any (P ∘ f) xs ↔ Any P (map f xs) @@ -271,8 +271,8 @@ module _ {P : Pred A p} where ++⁺∘++⁻ [] p = refl ++⁺∘++⁻ (x ∷ xs) (here p) = refl ++⁺∘++⁻ (x ∷ xs) (there p) with ++⁻ xs p | ++⁺∘++⁻ xs p - ++⁺∘++⁻ (x ∷ xs) (there p) | inj₁ p′ | ih = P.cong there ih - ++⁺∘++⁻ (x ∷ xs) (there p) | inj₂ p′ | ih = P.cong there ih + ++⁺∘++⁻ (x ∷ xs) (there p) | inj₁ p′ | ih = ≡.cong there ih + ++⁺∘++⁻ (x ∷ xs) (there p) | inj₂ p′ | ih = ≡.cong there ih ++⁻∘++⁺ : ∀ {n m} (xs : Vec A n) {ys : Vec A m} (p : Any P xs ⊎ Any P ys) → ++⁻ xs ([ ++⁺ˡ , ++⁺ʳ xs ]′ p) ≡ p @@ -336,10 +336,10 @@ module _ {P : Pred A p} where concat⁺ (concat⁻ xss p) ≡ p concat⁺∘concat⁻ (xs ∷ xss) p with ++⁻ xs p in eq ... | inj₁ pxs - = P.trans (P.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (P.sym eq)) + = ≡.trans (≡.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (≡.sym eq)) $ ++⁺∘++⁻ xs p ... | inj₂ pxss rewrite concat⁺∘concat⁻ xss pxss - = P.trans (P.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (P.sym eq)) + = ≡.trans (≡.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (≡.sym eq)) $ ++⁺∘++⁻ xs p concat⁻∘concat⁺ : ∀ {n m} {xss : Vec (Vec A n) m} (p : Any (Any P) xss) → @@ -402,7 +402,7 @@ module _ {P : Pred B p} where mapWith∈⁺ f (mapWith∈⁻ xs f p) ≡ p to∘from (y ∷ xs) f (here p) = refl to∘from (y ∷ xs) f (there p) = - P.cong there $ to∘from xs (f ∘ there) p + ≡.cong there $ to∘from xs (f ∘ there) p ------------------------------------------------------------------------ -- _∷_ diff --git a/src/Data/Vec/Relation/Unary/Unique/Setoid/Properties.agda b/src/Data/Vec/Relation/Unary/Unique/Setoid/Properties.agda index 692f2530bf..fa19dc3e21 100644 --- a/src/Data/Vec/Relation/Unary/Unique/Setoid/Properties.agda +++ b/src/Data/Vec/Relation/Unary/Unique/Setoid/Properties.agda @@ -20,7 +20,7 @@ open import Function.Base using (_∘_; id) open import Level using (Level) open import Relation.Binary.Core using (Rel) open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary.Negation using (contradiction; contraposition) private @@ -71,7 +71,7 @@ module _ (S : Setoid a ℓ) where open Setoid S lookup-injective : ∀ {n xs} → Unique S {n} xs → ∀ i j → lookup xs i ≈ lookup xs j → i ≡ j - lookup-injective (px ∷ pxs) zero zero eq = P.refl + lookup-injective (px ∷ pxs) zero zero eq = ≡.refl lookup-injective (px ∷ pxs) zero (suc j) eq = contradiction eq (All.lookup⁺ px j) lookup-injective (px ∷ pxs) (suc i) zero eq = contradiction (sym eq) (All.lookup⁺ px i) - lookup-injective (px ∷ pxs) (suc i) (suc j) eq = P.cong suc (lookup-injective pxs i j eq) + lookup-injective (px ∷ pxs) (suc i) (suc j) eq = ≡.cong suc (lookup-injective pxs i j eq) From a616ce6d71c295bf23a3b14a69bed2ae82c16322 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 14 Feb 2024 10:46:36 +0000 Subject: [PATCH 4/5] simplified imports; fixed `README` link --- src/Data/Fin/Substitution/Example.agda | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/Data/Fin/Substitution/Example.agda b/src/Data/Fin/Substitution/Example.agda index 5e5205ae14..7ac3061b11 100644 --- a/src/Data/Fin/Substitution/Example.agda +++ b/src/Data/Fin/Substitution/Example.agda @@ -12,7 +12,7 @@ module Data.Fin.Substitution.Example where {-# WARNING_ON_IMPORT "Data.Fin.Substitution.Example was deprecated in v2.0. -Please see README.Data.Nat.Fin.Substitution.UntypedLambda instead." +Please see README.Data.Fin.Substitution.UntypedLambda instead." #-} open import Data.Fin.Substitution @@ -20,9 +20,9 @@ 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 as ≡ - using (_≡_; refl; sym; cong; cong₂) -open ≡.≡-Reasoning +open import Relation.Binary.PropositionalEquality + using (_≡_; refl; sym; cong; cong₂; module ≡-Reasoning) +open ≡-Reasoning open import Relation.Binary.Construct.Closure.ReflexiveTransitive using (Star; ε; _◅_) From 3e4f8b1c5cf645d9bc00500c43be42f124f24c3f Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 14 Feb 2024 12:43:55 +0000 Subject: [PATCH 5/5] simplified imports --- src/Algebra/Solver/Monoid.agda | 4 +- src/Codata/Guarded/Stream/Properties.agda | 84 +++++++++---------- src/Codata/Musical/Colist.agda | 18 ++-- src/Codata/Musical/Colist/Infinite-merge.agda | 16 ++-- .../Colist/Relation/Unary/Any/Properties.agda | 12 +-- src/Data/Container/Combinator/Properties.agda | 24 +++--- src/Data/Container/Indexed.agda | 10 +-- src/Data/Container/Indexed/Combinator.agda | 12 +-- .../Container/Relation/Binary/Pointwise.agda | 3 +- .../Relation/Unary/Any/Properties.agda | 18 ++-- src/Data/Digit.agda | 6 +- src/Data/Fin/Permutation.agda | 12 +-- src/Data/Graph/Acyclic.agda | 18 ++-- .../Propositional/Properties/Core.agda | 12 +-- .../Binary/Equality/Propositional.agda | 8 +- .../Prefix/Heterogeneous/Properties.agda | 8 +- .../Binary/Sublist/Heterogeneous.agda | 1 - .../Suffix/Heterogeneous/Properties.agda | 8 +- .../List/Relation/Ternary/Interleaving.agda | 6 +- .../List/Relation/Unary/Any/Properties.agda | 68 +++++++-------- src/Data/Maybe/Relation/Unary/All.agda | 4 +- src/Data/Maybe/Relation/Unary/Any.agda | 2 +- src/Data/Nat/LCM.agda | 12 +-- .../Dependent/Propositional/WithK.agda | 1 - .../Relation/Binary/Pointwise/Dependent.agda | 2 +- src/Data/Sum/Relation/Binary/LeftOrder.agda | 2 +- .../Relation/Binary/Equality/Setoid.agda | 2 +- src/Data/Vec/Properties/WithK.agda | 4 +- src/Data/Vec/Recursive/Properties.agda | 33 ++++---- src/Data/Vec/Relation/Binary/Lex/Strict.agda | 10 +-- src/Data/Vec/Relation/Unary/All.agda | 8 +- .../Vec/Relation/Unary/Any/Properties.agda | 31 +++---- src/Effect/Applicative.agda | 2 +- src/Effect/Applicative/Indexed.agda | 4 +- src/Function/Endomorphism/Propositional.agda | 10 +-- src/Function/Related/TypeIsomorphisms.agda | 52 ++++++------ .../Construct/Add/Infimum/NonStrict.agda | 9 +- .../Binary/Construct/Add/Infimum/Strict.agda | 11 ++- .../Construct/Add/Supremum/NonStrict.agda | 10 +-- .../Binary/Construct/Add/Supremum/Strict.agda | 11 ++- .../Closure/Reflexive/Properties.agda | 2 +- .../Binary/Indexed/Heterogeneous/Bundles.agda | 3 - .../Binary/Indexed/Heterogeneous/Core.agda | 1 - .../Indexed/Heterogeneous/Definitions.agda | 2 - .../Binary/Indexed/Homogeneous/Bundles.agda | 7 +- src/Tactic/Cong.agda | 2 +- 46 files changed, 286 insertions(+), 299 deletions(-) diff --git a/src/Algebra/Solver/Monoid.agda b/src/Algebra/Solver/Monoid.agda index 2aea92797c..ba6aabee06 100644 --- a/src/Algebra/Solver/Monoid.agda +++ b/src/Algebra/Solver/Monoid.agda @@ -22,7 +22,7 @@ open import Data.Vec.Base using (Vec; lookup) open import Function.Base using (_∘_; _$_) open import Relation.Binary.Definitions using (Decidable) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; cong) import Relation.Binary.Reflection open import Relation.Nullary import Relation.Nullary.Decidable as Dec @@ -128,7 +128,7 @@ prove′ e₁ e₂ = lemma : normalise e₁ ≡ normalise e₂ → ∀ ρ → ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ lemma eq ρ = R.prove ρ e₁ e₂ (begin - ⟦ normalise e₁ ⟧⇓ ρ ≡⟨ ≡.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/Properties.agda b/src/Codata/Guarded/Stream/Properties.agda index 7a547e6ec2..c5f8e6b0cc 100644 --- a/src/Codata/Guarded/Stream/Properties.agda +++ b/src/Codata/Guarded/Stream/Properties.agda @@ -20,7 +20,7 @@ open import Data.Product.Base as Prod using (_×_; _,_; proj₁; proj₂) open import Data.Vec.Base as Vec using (Vec; _∷_) open import Function.Base using (const; flip; id; _∘′_; _$′_; _⟨_⟩_; _∘₂′_) open import Level using (Level) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; cong; cong₂) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong; cong₂) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) @@ -39,7 +39,7 @@ cong-lookup : ∀ {as bs : Stream A} → as ≈ bs → ∀ n → lookup as n ≡ cong-lookup = B.lookup⁺ cong-take : ∀ n {as bs : Stream A} → as ≈ bs → take n as ≡ take n bs -cong-take zero as≈bs = ≡.refl +cong-take zero as≈bs = refl cong-take (suc n) as≈bs = cong₂ _∷_ (as≈bs .head) (cong-take n (as≈bs .tail)) cong-drop : ∀ n {as bs : Stream A} → as ≈ bs → drop n as ≈ drop n bs @@ -63,7 +63,7 @@ cong-concat ass≈bss = cong-++-concat [] ass≈bss cong-++-concat : ∀ (as : List A) {ass bss} → ass ≈ bss → ++-concat as ass ≈ ++-concat as bss cong-++-concat [] ass≈bss .head = cong List⁺.head (ass≈bss .head) cong-++-concat [] ass≈bss .tail rewrite ass≈bss .head = cong-++-concat _ (ass≈bss .tail) - cong-++-concat (a ∷ as) ass≈bss .head = ≡.refl + cong-++-concat (a ∷ as) ass≈bss .head = refl cong-++-concat (a ∷ as) ass≈bss .tail = cong-++-concat as ass≈bss cong-interleave : {as bs cs ds : Stream A} → as ≈ bs → cs ≈ ds → @@ -79,11 +79,11 @@ cong-chunksOf n as≈bs .tail = cong-chunksOf n (cong-drop n as≈bs) -- Properties of repeat lookup-repeat : ∀ n (a : A) → lookup (repeat a) n ≡ a -lookup-repeat zero a = ≡.refl +lookup-repeat zero a = refl lookup-repeat (suc n) a = lookup-repeat n a splitAt-repeat : ∀ n (a : A) → splitAt n (repeat a) ≡ (Vec.replicate n a , repeat a) -splitAt-repeat zero a = ≡.refl +splitAt-repeat zero a = refl splitAt-repeat (suc n) a = cong (Prod.map₁ (a ∷_)) (splitAt-repeat n a) take-repeat : ∀ n (a : A) → take n (repeat a) ≡ Vec.replicate n a @@ -93,28 +93,28 @@ drop-repeat : ∀ n (a : A) → drop n (repeat a) ≡ repeat a drop-repeat n a = cong proj₂ (splitAt-repeat n a) map-repeat : ∀ (f : A → B) a → map f (repeat a) ≈ repeat (f a) -map-repeat f a .head = ≡.refl +map-repeat f a .head = refl map-repeat f a .tail = map-repeat f a ap-repeat : ∀ (f : A → B) a → ap (repeat f) (repeat a) ≈ repeat (f a) -ap-repeat f a .head = ≡.refl +ap-repeat f a .head = refl ap-repeat f a .tail = ap-repeat f a ap-repeatˡ : ∀ (f : A → B) as → ap (repeat f) as ≈ map f as -ap-repeatˡ f as .head = ≡.refl +ap-repeatˡ f as .head = refl ap-repeatˡ f as .tail = ap-repeatˡ f (as .tail) ap-repeatʳ : ∀ (fs : Stream (A → B)) a → ap fs (repeat a) ≈ map (_$′ a) fs -ap-repeatʳ fs a .head = ≡.refl +ap-repeatʳ fs a .head = refl ap-repeatʳ fs a .tail = ap-repeatʳ (fs .tail) a interleave-repeat : (a : A) → interleave (repeat a) (repeat a) ≈ repeat a -interleave-repeat a .head = ≡.refl +interleave-repeat a .head = refl interleave-repeat a .tail = interleave-repeat a zipWith-repeat : ∀ (f : A → B → C) a b → zipWith f (repeat a) (repeat b) ≈ repeat (f a b) -zipWith-repeat f a b .head = ≡.refl +zipWith-repeat f a b .head = refl zipWith-repeat f a b .tail = zipWith-repeat f a b chunksOf-repeat : ∀ n (a : A) → chunksOf n (repeat a) ≈ repeat (Vec.replicate n a) @@ -125,7 +125,7 @@ chunksOf-repeat n a = begin go where go : chunksOf n (repeat a) ≈∞ repeat (Vec.replicate n a) go .head = take-repeat n a go .tail = - chunksOf n (drop n (repeat a)) ≡⟨ ≡.cong (chunksOf n) (drop-repeat n a) ⟩ + chunksOf n (drop n (repeat a)) ≡⟨ cong (chunksOf n) (drop-repeat n a) ⟩ chunksOf n (repeat a) ↺⟨ go ⟩ repeat (Vec.replicate n a) ∎ @@ -133,34 +133,34 @@ chunksOf-repeat n a = begin go where -- Properties of map map-const : (a : A) (bs : Stream B) → map (const a) bs ≈ repeat a -map-const a bs .head = ≡.refl +map-const a bs .head = refl map-const a bs .tail = map-const a (bs .tail) map-id : (as : Stream A) → map id as ≈ as -map-id as .head = ≡.refl +map-id as .head = refl map-id as .tail = map-id (as .tail) map-∘ : ∀ (g : B → C) (f : A → B) as → map g (map f as) ≈ map (g ∘′ f) as -map-∘ g f as .head = ≡.refl +map-∘ g f as .head = refl map-∘ g f as .tail = map-∘ g f (as .tail) map-unfold : ∀ (g : B → C) (f : A → A × B) a → map g (unfold f a) ≈ unfold (Prod.map₂ g ∘′ f) a -map-unfold g f a .head = ≡.refl +map-unfold g f a .head = refl map-unfold g f a .tail = map-unfold g f (proj₁ (f a)) map-drop : ∀ (f : A → B) n as → map f (drop n as) ≡ drop n (map f as) -map-drop f zero as = ≡.refl +map-drop f zero as = refl map-drop f (suc n) as = map-drop f n (as .tail) map-zipWith : ∀ (g : C → D) (f : A → B → C) as bs → map g (zipWith f as bs) ≈ zipWith (g ∘₂′ f) as bs -map-zipWith g f as bs .head = ≡.refl +map-zipWith g f as bs .head = refl map-zipWith g f as bs .tail = map-zipWith g f (as .tail) (bs .tail) map-interleave : ∀ (f : A → B) as bs → map f (interleave as bs) ≈ interleave (map f as) (map f bs) -map-interleave f as bs .head = ≡.refl +map-interleave f as bs .head = refl map-interleave f as bs .tail = map-interleave f bs (as .tail) map-concat : ∀ (f : A → B) ass → map f (concat ass) ≈ concat (map (List⁺.map f) ass) @@ -168,9 +168,9 @@ map-concat f ass = map-++-concat [] ass where open Concat map-++-concat : ∀ acc ass → map f (++-concat acc ass) ≈ ++-concat (List.map f acc) (map (List⁺.map f) ass) - map-++-concat [] ass .head = ≡.refl + map-++-concat [] ass .head = refl map-++-concat [] ass .tail = map-++-concat (ass .head .List⁺.tail) (ass .tail) - map-++-concat (a ∷ as) ass .head = ≡.refl + map-++-concat (a ∷ as) ass .head = refl map-++-concat (a ∷ as) ass .tail = map-++-concat as ass map-cycle : ∀ (f : A → B) as → map f (cycle as) ≈ cycle (List⁺.map f as) @@ -186,29 +186,29 @@ map-cycle f as = run -- Properties of lookup lookup-drop : ∀ m (as : Stream A) n → lookup (drop m as) n ≡ lookup as (m + n) -lookup-drop zero as n = ≡.refl +lookup-drop zero as n = refl lookup-drop (suc m) as n = lookup-drop m (as .tail) n lookup-map : ∀ n (f : A → B) as → lookup (map f as) n ≡ f (lookup as n) -lookup-map zero f as = ≡.refl +lookup-map zero f as = refl lookup-map (suc n) f as = lookup-map n f (as . tail) lookup-iterate : ∀ n f (x : A) → lookup (iterate f x) n ≡ ℕ.iterate f x n -lookup-iterate zero f x = ≡.refl +lookup-iterate zero f x = refl lookup-iterate (suc n) f x = lookup-iterate n f (f x) lookup-zipWith : ∀ n (f : A → B → C) as bs → lookup (zipWith f as bs) n ≡ f (lookup as n) (lookup bs n) -lookup-zipWith zero f as bs = ≡.refl +lookup-zipWith zero f as bs = refl lookup-zipWith (suc n) f as bs = lookup-zipWith n f (as .tail) (bs .tail) lookup-unfold : ∀ n (f : A → A × B) a → lookup (unfold f a) n ≡ proj₂ (f (ℕ.iterate (proj₁ ∘′ f) a n)) -lookup-unfold zero f a = ≡.refl +lookup-unfold zero f a = refl lookup-unfold (suc n) f a = lookup-unfold n f (proj₁ (f a)) lookup-tabulate : ∀ n (f : ℕ → A) → lookup (tabulate f) n ≡ f n -lookup-tabulate zero f = ≡.refl +lookup-tabulate zero f = refl lookup-tabulate (suc n) f = lookup-tabulate n (f ∘′ suc) lookup-transpose : ∀ n (ass : List (Stream A)) → @@ -237,33 +237,33 @@ lookup-tails zero as = B.refl lookup-tails (suc n) as = lookup-tails n (as .tail) lookup-evens : ∀ n (as : Stream A) → lookup (evens as) n ≡ lookup as (n * 2) -lookup-evens zero as = ≡.refl +lookup-evens zero as = refl lookup-evens (suc n) as = lookup-evens n (as .tail .tail) lookup-odds : ∀ n (as : Stream A) → lookup (odds as) n ≡ lookup as (suc (n * 2)) -lookup-odds zero as = ≡.refl +lookup-odds zero as = refl lookup-odds (suc n) as = lookup-odds n (as .tail .tail) lookup-interleave-even : ∀ n (as bs : Stream A) → lookup (interleave as bs) (n * 2) ≡ lookup as n -lookup-interleave-even zero as bs = ≡.refl +lookup-interleave-even zero as bs = refl lookup-interleave-even (suc n) as bs = lookup-interleave-even n (as .tail) (bs .tail) lookup-interleave-odd : ∀ n (as bs : Stream A) → lookup (interleave as bs) (suc (n * 2)) ≡ lookup bs n -lookup-interleave-odd zero as bs = ≡.refl +lookup-interleave-odd zero as bs = refl lookup-interleave-odd (suc n) as bs = lookup-interleave-odd n (as .tail) (bs .tail) ------------------------------------------------------------------------ -- Properties of take take-iterate : ∀ n f (x : A) → take n (iterate f x) ≡ Vec.iterate f x n -take-iterate zero f x = ≡.refl +take-iterate zero f x = refl take-iterate (suc n) f x = cong (x ∷_) (take-iterate n f (f x)) take-zipWith : ∀ n (f : A → B → C) as bs → take n (zipWith f as bs) ≡ Vec.zipWith f (take n as) (take n bs) -take-zipWith zero f as bs = ≡.refl +take-zipWith zero f as bs = refl take-zipWith (suc n) f as bs = cong (f (as .head) (bs .head) ∷_) (take-zipWith n f (as .tail) (bs . tail)) @@ -271,21 +271,21 @@ take-zipWith (suc n) f as bs = -- Properties of drop drop-drop : ∀ m n (as : Stream A) → drop n (drop m as) ≡ drop (m + n) as -drop-drop zero n as = ≡.refl +drop-drop zero n as = refl drop-drop (suc m) n as = drop-drop m n (as .tail) drop-zipWith : ∀ n (f : A → B → C) as bs → drop n (zipWith f as bs) ≡ zipWith f (drop n as) (drop n bs) -drop-zipWith zero f as bs = ≡.refl +drop-zipWith zero f as bs = refl drop-zipWith (suc n) f as bs = drop-zipWith n f (as .tail) (bs .tail) drop-ap : ∀ n (fs : Stream (A → B)) as → drop n (ap fs as) ≡ ap (drop n fs) (drop n as) -drop-ap zero fs as = ≡.refl +drop-ap zero fs as = refl drop-ap (suc n) fs as = drop-ap n (fs .tail) (as .tail) drop-iterate : ∀ n f (x : A) → drop n (iterate f x) ≡ iterate f (ℕ.iterate f x n) -drop-iterate zero f x = ≡.refl +drop-iterate zero f x = refl drop-iterate (suc n) f x = drop-iterate n f (f x) ------------------------------------------------------------------------ @@ -293,25 +293,25 @@ drop-iterate (suc n) f x = drop-iterate n f (f x) zipWith-defn : ∀ (f : A → B → C) as bs → zipWith f as bs ≈ (repeat f ⟨ ap ⟩ as ⟨ ap ⟩ bs) -zipWith-defn f as bs .head = ≡.refl +zipWith-defn f as bs .head = refl zipWith-defn f as bs .tail = zipWith-defn f (as .tail) (bs .tail) zipWith-const : (as : Stream A) (bs : Stream B) → zipWith const as bs ≈ as -zipWith-const as bs .head = ≡.refl +zipWith-const as bs .head = refl zipWith-const as bs .tail = zipWith-const (as .tail) (bs .tail) zipWith-flip : ∀ (f : A → B → C) as bs → zipWith (flip f) as bs ≈ zipWith f bs as -zipWith-flip f as bs .head = ≡.refl +zipWith-flip f as bs .head = refl zipWith-flip f as bs .tail = zipWith-flip f (as .tail) (bs. tail) ------------------------------------------------------------------------ -- Properties of interleave interleave-evens-odds : (as : Stream A) → interleave (evens as) (odds as) ≈ as -interleave-evens-odds as .head = ≡.refl -interleave-evens-odds as .tail .head = ≡.refl +interleave-evens-odds as .head = refl +interleave-evens-odds as .tail .head = refl interleave-evens-odds as .tail .tail = interleave-evens-odds (as .tail .tail) ------------------------------------------------------------------------ diff --git a/src/Codata/Musical/Colist.agda b/src/Codata/Musical/Colist.agda index 78e8703be6..74bd2a2aa4 100644 --- a/src/Codata/Musical/Colist.agda +++ b/src/Codata/Musical/Colist.agda @@ -33,7 +33,7 @@ open import Relation.Binary.Definitions using (Transitive; Antisymmetric) import Relation.Binary.Construct.FromRel as Ind import Relation.Binary.Reasoning.Preorder as ≲-Reasoning import Relation.Binary.Reasoning.PartialOrder as ≤-Reasoning -open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) open import Relation.Binary.Reasoning.Syntax open import Relation.Nullary.Reflects using (invert) open import Relation.Nullary @@ -107,23 +107,23 @@ Any-∈ {P = P} = mk↔ₛ′ from∘to where to : ∀ {xs} → Any P xs → ∃ λ x → x ∈ xs × P x - to (here p) = _ , here ≡.refl , p + to (here p) = _ , here refl , 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 ≡.refl) p = here p + from (here refl) p = here p from (there x∈xs) p = there (from x∈xs p) to∘from : ∀ {x xs} (x∈xs : x ∈ xs) (p : P x) → to (from x∈xs p) ≡ (x , x∈xs , p) - to∘from (here ≡.refl) p = ≡.refl + to∘from (here refl) p = refl to∘from (there x∈xs) p = - ≡.cong (Product.map id (Product.map there id)) (to∘from x∈xs 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 - from∘to (here _) = ≡.refl - from∘to (there p) = ≡.cong there (from∘to p) + from∘to (here _) = refl + from∘to (there p) = cong there (from∘to p) -- Prefixes are subsets. @@ -220,7 +220,7 @@ infixr 5 _∷_ module Finite-injective where ∷-injective : ∀ {x : A} {xs p q} → (Finite (x ∷ xs) ∋ x ∷ p) ≡ x ∷ q → p ≡ q - ∷-injective ≡.refl = ≡.refl + ∷-injective refl = refl -- Infinite xs means that xs has infinite length. @@ -230,7 +230,7 @@ data Infinite {A : Set a} : Colist A → Set a where module Infinite-injective where ∷-injective : ∀ {x : A} {xs p q} → (Infinite (x ∷ xs) ∋ x ∷ p) ≡ x ∷ q → p ≡ q - ∷-injective ≡.refl = ≡.refl + ∷-injective refl = refl -- Colists which are not finite are infinite. diff --git a/src/Codata/Musical/Colist/Infinite-merge.agda b/src/Codata/Musical/Colist/Infinite-merge.agda index 69d07b9523..c4343bd9e5 100644 --- a/src/Codata/Musical/Colist/Infinite-merge.agda +++ b/src/Codata/Musical/Colist/Infinite-merge.agda @@ -25,7 +25,7 @@ open import Function.Related.TypeIsomorphisms open import Level open import Relation.Unary using (Pred) import Induction.WellFounded as WF -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) import Relation.Binary.Construct.On as On @@ -149,9 +149,9 @@ Any-merge {P = P} xss = mk↔ₛ′ (proj₁ ∘ to xss) from to∘from (proj₂ from-injective : ∀ {xss} (p₁ p₂ : Any Q xss) → from p₁ ≡ from p₂ → p₁ ≡ p₂ - from-injective (here (inj₁ p)) (here (inj₁ .p)) ≡.refl = ≡.refl + from-injective (here (inj₁ p)) (here (inj₁ .p)) refl = refl from-injective (here (inj₂ p₁)) (here (inj₂ p₂)) eq = - ≡.cong (here ∘ inj₂) $ + cong (here ∘ inj₂) $ inj₁-injective $ Injection.injective (↔⇒↣ (↔-sym (Any-⋎P _))) $ there-injective eq @@ -166,7 +166,7 @@ Any-merge {P = P} xss = mk↔ₛ′ (proj₁ ∘ to xss) from to∘from (proj₂ (there-injective eq) ... | () from-injective (there {x = _ , xs} p₁) (there p₂) eq = - ≡.cong there $ + cong there $ from-injective p₁ p₂ $ inj₂-injective $ Injection.injective (↔⇒↣ (↔-sym (Any-⋎P xs))) $ @@ -188,15 +188,15 @@ Any-merge {P = P} xss = mk↔ₛ′ (proj₁ ∘ to xss) from to∘from (proj₂ step : ∀ p → WF.WfRec (_<′_ on size) InputPred p → InputPred p step ([] , ()) rec - step ((x , xs) ∷ xss , here p) rec = here (inj₁ p) , ≡.refl + step ((x , xs) ∷ xss , here p) rec = here (inj₁ p) , refl step ((x , xs) ∷ xss , there p) rec with Inverse.to (Any-⋎P xs) p | Inverse.strictlyInverseʳ (Any-⋎P xs) p | index-Any-⋎P xs p - ... | inj₁ q | ≡.refl | _ = here (inj₂ q) , ≡.refl - ... | inj₂ q | ≡.refl | q≤p = + ... | inj₁ q | refl | _ = here (inj₂ q) , refl + ... | inj₂ q | refl | q≤p = Product.map there - (≡.cong (there ∘ (Inverse.from (Any-⋎P xs)) ∘ inj₂)) + (cong (there ∘ (Inverse.from (Any-⋎P xs)) ∘ inj₂)) (rec (s≤′s q≤p)) to∘from = λ p → from-injective _ _ (proj₂ (to xss (from p))) diff --git a/src/Codata/Musical/Colist/Relation/Unary/Any/Properties.agda b/src/Codata/Musical/Colist/Relation/Unary/Any/Properties.agda index 93c767f0b7..438922dc42 100644 --- a/src/Codata/Musical/Colist/Relation/Unary/Any/Properties.agda +++ b/src/Codata/Musical/Colist/Relation/Unary/Any/Properties.agda @@ -21,7 +21,7 @@ open import Function.Base using (_∋_; _∘_) open import Function.Bundles open import Level using (Level; _⊔_) open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality as ≡ +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) open import Relation.Unary using (Pred) @@ -66,15 +66,15 @@ Any-cong {A = A} {P} {Q} {xs} {ys} P↔Q xs≈ys = to∘from : ∀ {xs ys} (xs≈ys : xs ≈ ys) (q : Any Q ys) → to xs≈ys (from xs≈ys q) ≡ q - to∘from (x ∷ xs≈) (there q) = ≡.cong there (to∘from (♭ xs≈) q) + to∘from (x ∷ xs≈) (there q) = cong there (to∘from (♭ xs≈) q) to∘from (x ∷ xs≈) (here qx) = - ≡.cong here (Inverse.strictlyInverseˡ P↔Q qx) + cong here (Inverse.strictlyInverseˡ P↔Q qx) from∘to : ∀ {xs ys} (xs≈ys : xs ≈ ys) (p : Any P xs) → from xs≈ys (to xs≈ys p) ≡ p - from∘to (x ∷ xs≈) (there p) = ≡.cong there (from∘to (♭ xs≈) p) + from∘to (x ∷ xs≈) (there p) = cong there (from∘to (♭ xs≈) p) from∘to (x ∷ xs≈) (here px) = - ≡.cong here (Inverse.strictlyInverseʳ P↔Q px) + cong here (Inverse.strictlyInverseʳ P↔Q px) ------------------------------------------------------------------------ -- map @@ -164,7 +164,7 @@ lookup-index (there p) = lookup-index p index-Any-resp : ∀ {f : ∀ {x} → P x → Q x} {xs ys} (xs≈ys : xs ≈ ys) (p : Any P xs) → index (Any-resp f xs≈ys p) ≡ index p -index-Any-resp (x ∷ xs≈) (here px) = ≡.refl +index-Any-resp (x ∷ xs≈) (here px) = refl index-Any-resp (x ∷ xs≈) (there p) = cong suc (index-Any-resp (♭ xs≈) p) diff --git a/src/Data/Container/Combinator/Properties.agda b/src/Data/Container/Combinator/Properties.agda index 5854138ec1..f4b9568eb2 100644 --- a/src/Data/Container/Combinator/Properties.agda +++ b/src/Data/Container/Combinator/Properties.agda @@ -18,7 +18,7 @@ open import Data.Sum.Base as S using (inj₁; inj₂; [_,_]′; [_,_]) open import Function.Base as F using (_∘′_) open import Function.Bundles open import Level using (_⊔_; lower) -open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_) +open import Relation.Binary.PropositionalEquality using (_≡_; _≗_; refl; cong) -- I have proved some of the correctness statements under the -- assumption of functional extensionality. I could have reformulated @@ -27,35 +27,35 @@ open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_) module Identity where correct : ∀ {s p x} {X : Set x} → ⟦ id {s} {p} ⟧ X ↔ F.id X - correct {X = X} = mk↔ₛ′ from-id to-id (λ _ → ≡.refl) (λ _ → ≡.refl) + correct {X = X} = mk↔ₛ′ from-id to-id (λ _ → refl) (λ _ → refl) module Constant (ext : ∀ {ℓ ℓ′} → Extensionality ℓ ℓ′) where correct : ∀ {x p y} (X : Set x) {Y : Set y} → ⟦ const {x} {p ⊔ y} X ⟧ Y ↔ F.const X Y - correct {x} {y} X {Y} = mk↔ₛ′ (from-const X) (to-const X) (λ _ → ≡.refl) from∘to + correct {x} {y} X {Y} = mk↔ₛ′ (from-const X) (to-const X) (λ _ → refl) from∘to where from∘to : (x : ⟦ const X ⟧ Y) → to-const X (proj₁ x) ≡ x - from∘to xs = ≡.cong (proj₁ xs ,_) (ext (λ x → ⊥-elim (lower x))) + from∘to xs = cong (proj₁ xs ,_) (ext (λ x → ⊥-elim (lower x))) module Composition {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) where correct : ∀ {x} {X : Set x} → ⟦ C₁ ∘ C₂ ⟧ X ↔ (⟦ C₁ ⟧ F.∘ ⟦ C₂ ⟧) X - correct {X = X} = mk↔ₛ′ (from-∘ C₁ C₂) (to-∘ C₁ C₂) (λ _ → ≡.refl) (λ _ → ≡.refl) + correct {X = X} = mk↔ₛ′ (from-∘ C₁ C₂) (to-∘ C₁ C₂) (λ _ → refl) (λ _ → refl) module Product (ext : ∀ {ℓ ℓ′} → Extensionality ℓ ℓ′) {s₁ s₂ p₁ p₂} (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) where correct : ∀ {x} {X : Set x} → ⟦ C₁ × C₂ ⟧ X ↔ (⟦ C₁ ⟧ X Prod.× ⟦ C₂ ⟧ X) - correct {X = X} = mk↔ₛ′ (from-× C₁ C₂) (to-× C₁ C₂) (λ _ → ≡.refl) from∘to + correct {X = X} = mk↔ₛ′ (from-× C₁ C₂) (to-× C₁ C₂) (λ _ → refl) from∘to where from∘to : (to-× C₁ C₂) F.∘ (from-× C₁ C₂) ≗ F.id from∘to (s , f) = - ≡.cong (s ,_) (ext [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ]) + cong (s ,_) (ext [ (λ _ → refl) , (λ _ → refl) ]) module IndexedProduct {i s p} {I : Set i} (Cᵢ : I → Container s p) where correct : ∀ {x} {X : Set x} → ⟦ Π I Cᵢ ⟧ X ↔ (∀ i → ⟦ Cᵢ i ⟧ X) - correct {X = X} = mk↔ₛ′ (from-Π I Cᵢ) (to-Π I Cᵢ) (λ _ → ≡.refl) (λ _ → ≡.refl) + correct {X = X} = mk↔ₛ′ (from-Π I Cᵢ) (to-Π I Cᵢ) (λ _ → refl) (λ _ → refl) module Sum {s₁ s₂ p} (C₁ : Container s₁ p) (C₂ : Container s₂ p) where @@ -63,16 +63,16 @@ module Sum {s₁ s₂ p} (C₁ : Container s₁ p) (C₂ : Container s₂ p) whe correct {X = X} = mk↔ₛ′ (from-⊎ C₁ C₂) (to-⊎ C₁ C₂) to∘from from∘to where from∘to : (to-⊎ C₁ C₂) F.∘ (from-⊎ C₁ C₂) ≗ F.id - from∘to (inj₁ s₁ , f) = ≡.refl - from∘to (inj₂ s₂ , f) = ≡.refl + from∘to (inj₁ s₁ , f) = refl + from∘to (inj₂ s₂ , f) = refl to∘from : (from-⊎ C₁ C₂) F.∘ (to-⊎ C₁ C₂) ≗ F.id - to∘from = [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] + to∘from = [ (λ _ → refl) , (λ _ → refl) ] module IndexedSum {i s p} {I : Set i} (C : I → Container s p) where correct : ∀ {x} {X : Set x} → ⟦ Σ I C ⟧ X ↔ (∃ λ i → ⟦ C i ⟧ X) - correct {X = X} = mk↔ₛ′ (from-Σ I C) (to-Σ I C) (λ _ → ≡.refl) (λ _ → ≡.refl) + correct {X = X} = mk↔ₛ′ (from-Σ I C) (to-Σ I C) (λ _ → refl) (λ _ → refl) module ConstantExponentiation {i s p} {I : Set i} (C : Container s p) where diff --git a/src/Data/Container/Indexed.agda b/src/Data/Container/Indexed.agda index 8507a2af39..b34657a4be 100644 --- a/src/Data/Container/Indexed.agda +++ b/src/Data/Container/Indexed.agda @@ -18,7 +18,7 @@ open import Function.Base renaming (id to ⟨id⟩; _∘_ to _⟨∘⟩_) open import Function using (_↔_; Inverse) open import Relation.Unary using (Pred; _⊆_) open import Relation.Binary.Core using (Rel; REL) -open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_; refl) +open import Relation.Binary.PropositionalEquality using (_≡_; _≗_; refl; trans; subst) ------------------------------------------------------------------------ @@ -98,7 +98,7 @@ module _ {i₁ i₂ o₁ o₂} Container I₁ O₁ c₁ r → (I₁ → I₂) → (O₁ → O₂) → Container I₂ O₂ c₂ r → Set _ C₁ ⇒C[ f / g ] C₂ = ContainerMorphism C₁ C₂ f g _≡_ (λ R₂ R₁ → R₂ ≡ R₁) - (λ r₂≡r₁ r₂ → ≡.subst ⟨id⟩ r₂≡r₁ r₂) + (λ r₂≡r₁ r₂ → subst ⟨id⟩ r₂≡r₁ r₂) -- Degenerate cases where no reindexing is performed. @@ -123,7 +123,7 @@ module _ {i o c r} {I : Set i} {O : Set o} where ⟪_⟫ : ∀ {i o c r ℓ} {I : Set i} {O : Set o} {C₁ C₂ : Container I O c r} → C₁ ⇒ C₂ → (X : Pred I ℓ) → ⟦ C₁ ⟧ X ⊆ ⟦ C₂ ⟧ X ⟪ m ⟫ X (c , k) = command m c , λ r₂ → - ≡.subst X (coherent m) (k (response m r₂)) + subst X (coherent m) (k (response m r₂)) module PlainMorphism {i o c r} {I : Set i} {O : Set o} where @@ -145,7 +145,7 @@ module PlainMorphism {i o c r} {I : Set i} {O : Set o} where f ∘ g = record { command = command f ⟨∘⟩ command g ; response = response g ⟨∘⟩ response f - ; coherent = coherent g ⟨ ≡.trans ⟩ coherent f + ; coherent = coherent g ⟨ trans ⟩ coherent f } -- Identity commutes with ⟪_⟫. @@ -187,7 +187,7 @@ module CartesianMorphism morphism : C₁ ⇒ C₂ morphism = record { command = command m - ; response = ≡.subst ⟨id⟩ (response m) + ; response = subst ⟨id⟩ (response m) ; coherent = coherent m } diff --git a/src/Data/Container/Indexed/Combinator.agda b/src/Data/Container/Indexed/Combinator.agda index cb17210281..825d173db7 100644 --- a/src/Data/Container/Indexed/Combinator.agda +++ b/src/Data/Container/Indexed/Combinator.agda @@ -21,8 +21,8 @@ open import Function.Indexed.Bundles using (_↔ᵢ_) open import Level open import Relation.Unary using (Pred; _⊆_; _∪_; _∩_; ⋃; ⋂) renaming (_⟨×⟩_ to _⟪×⟫_; _⟨⊙⟩_ to _⟪⊙⟫_; _⟨⊎⟩_ to _⟪⊎⟫_) -open import Relation.Binary.PropositionalEquality as ≡ - using (_≗_; refl) +open import Relation.Binary.PropositionalEquality + using (_≗_; refl; cong) private variable @@ -167,7 +167,7 @@ module Constant (ext : ∀ {ℓ} → Extensionality ℓ ℓ) where from = < F.id , F.const ⊥-elim > to∘from : _ - to∘from xs = ≡.cong (proj₁ xs ,_) (ext ⊥-elim) + to∘from xs = cong (proj₁ xs ,_) (ext ⊥-elim) module Duality where @@ -202,7 +202,7 @@ module Product (ext : ∀ {ℓ} → Extensionality ℓ ℓ) where from∘to : from ⟨∘⟩ to ≗ F.id from∘to (c , _) = - ≡.cong (c ,_) (ext [ (λ _ → refl) , (λ _ → refl) ]) + cong (c ,_) (ext [ (λ _ → refl) , (λ _ → refl) ]) module IndexedProduct where @@ -231,8 +231,8 @@ module Sum (ext : ∀ {ℓ₁ ℓ₂} → Extensionality ℓ₁ ℓ₂) where from (inj₂ (c , f)) = inj₂ c , λ{ (All.inj₂ r) → f r} from∘to : from ⟨∘⟩ to ≗ F.id - from∘to (inj₁ _ , _) = ≡.cong (inj₁ _ ,_) (ext λ{ (All.inj₁ r) → refl}) - from∘to (inj₂ _ , _) = ≡.cong (inj₂ _ ,_) (ext λ{ (All.inj₂ r) → refl}) + from∘to (inj₁ _ , _) = cong (inj₁ _ ,_) (ext λ{ (All.inj₁ r) → refl}) + from∘to (inj₂ _ , _) = cong (inj₂ _ ,_) (ext λ{ (All.inj₂ r) → refl}) to∘from : to ⟨∘⟩ from ≗ F.id to∘from = [ (λ _ → refl) , (λ _ → refl) ] diff --git a/src/Data/Container/Relation/Binary/Pointwise.agda b/src/Data/Container/Relation/Binary/Pointwise.agda index bc66fe077d..0e2df66f78 100644 --- a/src/Data/Container/Relation/Binary/Pointwise.agda +++ b/src/Data/Container/Relation/Binary/Pointwise.agda @@ -12,8 +12,7 @@ open import Data.Product.Base using (_,_; Σ-syntax; -,_; proj₁; proj₂) open import Function.Base using (_∘_) open import Level using (_⊔_) open import Relation.Binary.Core using (REL; _⇒_) -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_; subst; cong) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; subst) open import Data.Container.Core using (Container; ⟦_⟧) diff --git a/src/Data/Container/Relation/Unary/Any/Properties.agda b/src/Data/Container/Relation/Unary/Any/Properties.agda index 42ddeb6696..601d922a56 100644 --- a/src/Data/Container/Relation/Unary/Any/Properties.agda +++ b/src/Data/Container/Relation/Unary/Any/Properties.agda @@ -42,21 +42,21 @@ module _ {s p} (C : Container s p) {x} {X : Set x} {ℓ} {P : Pred X ℓ} where -- ◇ can be unwrapped to reveal the Σ type ↔Σ : ∀ {xs : ⟦ C ⟧ X} → ◇ C P xs ↔ ∃ λ p → P (proj₂ xs p) - ↔Σ {xs} = mk↔ₛ′ ◇.proof any (λ _ → ≡.refl) (λ _ → ≡.refl) + ↔Σ {xs} = mk↔ₛ′ ◇.proof any (λ _ → refl) (λ _ → refl) -- ◇ can be expressed using _∈_. ↔∈ : ∀ {xs : ⟦ C ⟧ X} → ◇ C P xs ↔ (∃ λ x → x ∈ xs × P x) - ↔∈ {xs} = mk↔ₛ′ to from to∘from (λ _ → ≡.refl) where + ↔∈ {xs} = mk↔ₛ′ to from to∘from (λ _ → refl) where to : ◇ C P xs → ∃ λ x → x ∈ xs × P x - to (any (p , Px)) = (proj₂ xs p , (any (p , ≡.refl)) , Px) + to (any (p , Px)) = (proj₂ xs p , (any (p , refl)) , Px) from : (∃ λ x → x ∈ xs × P x) → ◇ C P xs from (.(proj₂ xs p) , (any (p , refl)) , Px) = any (p , Px) to∘from : to ∘ from ≗ id - to∘from (.(proj₂ xs p) , any (p , refl) , Px) = ≡.refl + to∘from (.(proj₂ xs p) , any (p , refl) , Px) = refl module _ {s p} {C : Container s p} {x} {X : Set x} {ℓ₁ ℓ₂} {P₁ : Pred X ℓ₁} {P₂ : Pred X ℓ₂} where @@ -105,7 +105,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s flatten : ∀ (xss : ⟦ C₁ ⟧ (⟦ C₂ ⟧ X)) → ◇ C₁ (◇ C₂ P) xss ↔ ◇ (C₁ C.∘ C₂) P (Inverse.from (Composition.correct C₁ C₂) xss) - flatten xss = mk↔ₛ′ t f (λ _ → ≡.refl) (λ _ → ≡.refl) where + flatten xss = mk↔ₛ′ t f (λ _ → refl) (λ _ → refl) where ◇₁ = ◇ C₁; ◇₂ = ◇ C₂; ◇₁₂ = ◇ (C₁ C.∘ C₂) open Inverse @@ -132,11 +132,11 @@ module _ {s p} {C : Container s p} {x} {X : Set x} from = [ Any.map₂ inj₁ , Any.map₂ inj₂ ] from∘to : from ∘ to ≗ id - from∘to (any (pos , inj₁ p)) = ≡.refl - from∘to (any (pos , inj₂ q)) = ≡.refl + from∘to (any (pos , inj₁ p)) = refl + from∘to (any (pos , inj₂ q)) = refl to∘from : to ∘ from ≗ id - to∘from = [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] + to∘from = [ (λ _ → refl) , (λ _ → refl) ] -- Products "commute" with ◇. @@ -145,7 +145,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s ×◇↔◇◇× : ∀ {xs : ⟦ C₁ ⟧ X} {ys : ⟦ C₂ ⟧ Y} → ◇ C₁ (λ x → ◇ C₂ (λ y → P x × Q y) ys) xs ↔ (◇ C₁ P xs × ◇ C₂ Q ys) - ×◇↔◇◇× {xs} {ys} = mk↔ₛ′ to from (λ _ → ≡.refl) (λ _ → ≡.refl) + ×◇↔◇◇× {xs} {ys} = mk↔ₛ′ to from (λ _ → refl) (λ _ → refl) where ◇₁ = ◇ C₁; ◇₂ = ◇ C₂ diff --git a/src/Data/Digit.agda b/src/Data/Digit.agda index 82582aa472..49370776da 100644 --- a/src/Data/Digit.agda +++ b/src/Data/Digit.agda @@ -21,7 +21,7 @@ open import Data.Nat.DivMod open import Data.Nat.Induction open import Relation.Nullary.Decidable using (True; does; toWitness) open import Relation.Binary.Definitions using (Decidable) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; sym; cong) open import Function.Base using (_$_) ------------------------------------------------------------------------ @@ -85,7 +85,7 @@ toDigits base@(suc (suc k)) n = <′-rec Pred helper n Pred = λ n → ∃ λ ds → fromDigits ds ≡ n cons : ∀ {m} (r : Digit base) → Pred m → Pred (toℕ r + m * base) - cons r (ds , eq) = (r ∷ ds , ≡.cong (λ i → toℕ r + i * base) eq) + cons r (ds , eq) = (r ∷ ds , cong (λ i → toℕ r + i * base) eq) open ≤-Reasoning open +-*-Solver @@ -104,7 +104,7 @@ toDigits base@(suc (suc k)) n = <′-rec Pred helper n helper : ∀ n → <′-Rec Pred n → Pred n helper n rec with n divMod base - ... | result zero r eq = ([ r ] , ≡.sym eq) + ... | result zero r eq = ([ r ] , sym eq) ... | result (suc x) r refl = cons r (rec (lem x k (toℕ r))) ------------------------------------------------------------------------ diff --git a/src/Data/Fin/Permutation.agda b/src/Data/Fin/Permutation.agda index 96a557a4c7..97412c094d 100644 --- a/src/Data/Fin/Permutation.agda +++ b/src/Data/Fin/Permutation.agda @@ -28,9 +28,9 @@ open import Relation.Binary.Core using (Rel) open import Relation.Nullary using (does; ¬_; yes; no) open import Relation.Nullary.Decidable using (dec-yes; dec-no) open import Relation.Nullary.Negation using (contradiction) -open import Relation.Binary.PropositionalEquality as ≡ - using (_≡_; _≢_; refl; trans; sym; →-to-⟶; cong; cong₂) -open ≡.≡-Reasoning +open import Relation.Binary.PropositionalEquality + using (_≡_; _≢_; refl; sym; trans; subst; →-to-⟶; cong; cong₂; module ≡-Reasoning) +open ≡-Reasoning private variable @@ -67,10 +67,10 @@ _⟨$⟩ˡ_ : Permutation m n → Fin n → Fin m _⟨$⟩ˡ_ = Inverse.from inverseˡ : ∀ (π : Permutation m n) {i} → π ⟨$⟩ˡ (π ⟨$⟩ʳ i) ≡ i -inverseˡ π = Inverse.inverseʳ π ≡.refl +inverseˡ π = Inverse.inverseʳ π refl inverseʳ : ∀ (π : Permutation m n) {i} → π ⟨$⟩ʳ (π ⟨$⟩ˡ i) ≡ i -inverseʳ π = Inverse.inverseˡ π ≡.refl +inverseʳ π = Inverse.inverseˡ π refl ------------------------------------------------------------------------ -- Equality @@ -251,7 +251,7 @@ module _ (π : Permutation (suc m) (suc n)) where ↔⇒≡ {suc m} {suc n} π = cong suc (↔⇒≡ (remove 0F π)) fromPermutation : Permutation m n → Permutation′ m -fromPermutation π = ≡.subst (Permutation _) (sym (↔⇒≡ π)) π +fromPermutation π = subst (Permutation _) (sym (↔⇒≡ π)) π refute : m ≢ n → ¬ Permutation m n refute m≢n π = contradiction (↔⇒≡ π) m≢n diff --git a/src/Data/Graph/Acyclic.agda b/src/Data/Graph/Acyclic.agda index 26af8a3fd9..bf4c7267b7 100644 --- a/src/Data/Graph/Acyclic.agda +++ b/src/Data/Graph/Acyclic.agda @@ -27,7 +27,7 @@ open import Data.Vec.Base as Vec using (Vec; []; _∷_) open import Data.List.Base as List using (List; []; _∷_) open import Function.Base using (_$_; _∘′_; _∘_; id) open import Relation.Nullary -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) ------------------------------------------------------------------------ -- A lemma @@ -186,7 +186,7 @@ private test-nodes : nodes example ≡ (# 0 , 0) ∷ (# 1 , 1) ∷ (# 2 , 2) ∷ (# 3 , 3) ∷ (# 4 , 4) ∷ [] - test-nodes = ≡.refl + test-nodes = refl module _ {ℓ e} {N : Set ℓ} {E : Set e} where @@ -213,7 +213,7 @@ private test-edges : edges example ≡ (# 1 , 10 , # 1) ∷ (# 1 , 11 , # 1) ∷ (# 2 , 12 , # 0) ∷ [] - test-edges = ≡.refl + test-edges = refl -- The successors of a given node i (edge label × node number relative -- to i). @@ -225,7 +225,7 @@ sucs g i = successors $ head (g [ i ]) private test-sucs : sucs example (# 1) ≡ (10 , # 1) ∷ (11 , # 1) ∷ [] - test-sucs = ≡.refl + test-sucs = refl -- The predecessors of a given node i (node number relative to i × -- edge label). @@ -238,13 +238,13 @@ preds (c & g) (suc i) = (List.map (Prod.map suc id) $ preds g i) where p : ∀ {e} {E : Set e} {n} (i : Fin n) → E × Fin n → Maybe (Fin′ (suc i) × E) - p i (e , j) = Maybe.map (λ{ ≡.refl → zero , e }) (decToMaybe (i ≟ j)) + p i (e , j) = Maybe.map (λ{ refl → zero , e }) (decToMaybe (i ≟ j)) private test-preds : preds example (# 3) ≡ (# 1 , 10) ∷ (# 1 , 11) ∷ (# 2 , 12) ∷ [] - test-preds = ≡.refl + test-preds = refl ------------------------------------------------------------------------ -- Operations @@ -272,7 +272,7 @@ private context (# 3 , 3) [] & context (# 4 , 4) [] & ∅) - test-number = ≡.refl + test-number = refl -- Reverses all the edges in the graph. @@ -290,7 +290,7 @@ reverse {N = N} {E} g = private test-reverse : reverse (reverse example) ≡ example - test-reverse = ≡.refl + test-reverse = refl ------------------------------------------------------------------------ -- Views @@ -330,4 +330,4 @@ private node 3 [] ∷ node 4 [] ∷ [] - test-toForest = ≡.refl + test-toForest = refl diff --git a/src/Data/List/Membership/Propositional/Properties/Core.agda b/src/Data/List/Membership/Propositional/Properties/Core.agda index 568b567b37..7ad4f94573 100644 --- a/src/Data/List/Membership/Propositional/Properties/Core.agda +++ b/src/Data/List/Membership/Propositional/Properties/Core.agda @@ -20,8 +20,8 @@ open import Data.List.Membership.Propositional open import Data.Product.Base as Product using (_,_; proj₁; proj₂; uncurry′; ∃; _×_) open import Level using (Level) -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core + using (_≡_; refl; cong; subst) open import Relation.Unary using (Pred; _⊆_) private @@ -37,8 +37,8 @@ map∘find : ∀ {P : Pred A p} {xs} {f : _≡_ (proj₁ p′) ⊆ P} → f refl ≡ proj₂ (proj₂ p′) → Any.map f (proj₁ (proj₂ p′)) ≡ p -map∘find (here p) hyp = ≡.cong here hyp -map∘find (there p) hyp = ≡.cong there (map∘find p hyp) +map∘find (here p) hyp = cong here hyp +map∘find (there p) hyp = 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) → @@ -61,13 +61,13 @@ find-∈ (there x∈xs) rewrite find-∈ x∈xs = refl lose∘find : ∀ {P : Pred A p} {xs : List A} (p : Any P xs) → uncurry′ lose (proj₂ (find p)) ≡ p -lose∘find p = map∘find p ≡.refl +lose∘find p = map∘find p refl find∘lose : ∀ (P : Pred A p) {x xs} (x∈xs : x ∈ xs) (pp : P x) → find {P = P} (lose x∈xs pp) ≡ (x , x∈xs , pp) find∘lose P x∈xs p - rewrite find∘map x∈xs (flip (≡.subst P) p) + rewrite find∘map x∈xs (flip (subst P) p) | find-∈ x∈xs = refl diff --git a/src/Data/List/Relation/Binary/Equality/Propositional.agda b/src/Data/List/Relation/Binary/Equality/Propositional.agda index 2b458c29b9..ffae515086 100644 --- a/src/Data/List/Relation/Binary/Equality/Propositional.agda +++ b/src/Data/List/Relation/Binary/Equality/Propositional.agda @@ -16,7 +16,7 @@ module Data.List.Relation.Binary.Equality.Propositional {a} {A : Set a} where open import Data.List.Base import Data.List.Relation.Binary.Equality.Setoid as SetoidEquality -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) import Relation.Binary.PropositionalEquality.Properties as ≡ ------------------------------------------------------------------------ @@ -28,8 +28,8 @@ open SetoidEquality (≡.setoid A) public -- ≋ is propositional ≋⇒≡ : _≋_ ⇒ _≡_ -≋⇒≡ [] = ≡.refl -≋⇒≡ (≡.refl ∷ xs≈ys) = ≡.cong (_ ∷_) (≋⇒≡ xs≈ys) +≋⇒≡ [] = refl +≋⇒≡ (refl ∷ xs≈ys) = cong (_ ∷_) (≋⇒≡ xs≈ys) ≡⇒≋ : _≡_ ⇒ _≋_ -≡⇒≋ ≡.refl = ≋-refl +≡⇒≋ refl = ≋-refl diff --git a/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda index 8186fc9879..18f9871f1b 100644 --- a/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Prefix/Heterogeneous/Properties.agda @@ -28,7 +28,7 @@ open import Relation.Unary as U using (Pred) open import Relation.Binary.Core using (Rel; REL; _⇒_) open import Relation.Binary.Definitions using (Trans; Antisym; Irrelevant; Decidable) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; _≢_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; _≢_; refl; cong₂) private variable @@ -158,7 +158,7 @@ replicate⁺ (s≤s m≤n) r = r ∷ replicate⁺ m≤n r replicate⁻ : ∀ {m n a b} → m ≢ 0 → Prefix R (replicate m a) (replicate n b) → R a b -replicate⁻ {m = zero} {n} m≢0 r = ⊥-elim (m≢0 ≡.refl) +replicate⁻ {m = zero} {n} m≢0 r = ⊥-elim (m≢0 refl) replicate⁻ {m = suc m} {suc n} m≢0 rs = Prefix.head rs ------------------------------------------------------------------------ @@ -208,9 +208,9 @@ module _ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where irrelevant : Irrelevant R → Irrelevant (Prefix R) - irrelevant R-irr [] [] = ≡.refl + irrelevant R-irr [] [] = refl irrelevant R-irr (r ∷ rs) (r′ ∷ rs′) = - ≡.cong₂ _∷_ (R-irr r r′) (irrelevant R-irr rs rs′) + cong₂ _∷_ (R-irr r r′) (irrelevant R-irr rs rs′) ------------------------------------------------------------------------ -- Decidability diff --git a/src/Data/List/Relation/Binary/Sublist/Heterogeneous.agda b/src/Data/List/Relation/Binary/Sublist/Heterogeneous.agda index f66a29ee5e..a0fe6b150c 100644 --- a/src/Data/List/Relation/Binary/Sublist/Heterogeneous.agda +++ b/src/Data/List/Relation/Binary/Sublist/Heterogeneous.agda @@ -13,7 +13,6 @@ open import Data.List.Relation.Unary.Any using (Any; here; there) open import Level using (_⊔_) open import Relation.Binary.Core using (REL; _⇒_) open import Relation.Binary.Definitions using (_⟶_Respects_; Min) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Unary using (Pred) module Data.List.Relation.Binary.Sublist.Heterogeneous diff --git a/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda b/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda index 2090948734..abacb301d0 100644 --- a/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Suffix/Heterogeneous/Properties.agda @@ -27,8 +27,8 @@ open import Relation.Nullary.Negation using (contradiction) open import Relation.Binary.Core using (REL; Rel; _⇒_) open import Relation.Binary.Definitions as B using (Trans; Antisym; Irrelevant) -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_; _≢_; refl; sym; subst; subst₂) +open import Relation.Binary.PropositionalEquality.Core + using (_≡_; _≢_; refl; sym; cong; subst; subst₂) import Data.List.Properties as List import Data.List.Relation.Binary.Prefix.Heterogeneous.Properties as Prefix @@ -199,10 +199,10 @@ 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 irrelevant : Irrelevant R → Irrelevant (Suffix R) - irrelevant irr (here rs) (here rs₁) = ≡.cong here $ Pw.irrelevant irr rs rs₁ + irrelevant irr (here rs) (here rs₁) = cong here $ Pw.irrelevant irr rs rs₁ irrelevant irr (here rs) (there rsuf) = contradiction (Pointwise-length rs) (S[as][bs]⇒∣as∣≢1+∣bs∣ rsuf) irrelevant irr (there rsuf) (here rs) = contradiction (Pointwise-length rs) (S[as][bs]⇒∣as∣≢1+∣bs∣ rsuf) - irrelevant irr (there rsuf) (there rsuf₁) = ≡.cong there $ irrelevant irr rsuf rsuf₁ + irrelevant irr (there rsuf) (there rsuf₁) = cong there $ irrelevant irr rsuf rsuf₁ ------------------------------------------------------------------------ -- Decidability diff --git a/src/Data/List/Relation/Ternary/Interleaving.agda b/src/Data/List/Relation/Ternary/Interleaving.agda index b8201e1072..38ecbf5579 100644 --- a/src/Data/List/Relation/Ternary/Interleaving.agda +++ b/src/Data/List/Relation/Ternary/Interleaving.agda @@ -16,7 +16,7 @@ open import Data.Product.Base as Product using (∃; ∃₂; _×_; uncurry; _,_; open import Data.Sum.Base using (_⊎_; inj₁; inj₂) open import Function.Base open import Relation.Binary.Core using (REL; _⇒_) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) ------------------------------------------------------------------------ -- Definition @@ -60,9 +60,9 @@ module _ {a b c l r} {A : Set a} {B : Set b} {C : Set c} Interleaving _≡_ _≡_ csl csr cs × Pointwise L l csl × Pointwise R r csr break [] = -, [] , [] , [] break (l ∷ˡ sp) = let (_ , eq , pwl , pwr) = break sp in - -, ≡.refl ∷ˡ eq , l ∷ pwl , pwr + -, refl ∷ˡ eq , l ∷ pwl , pwr break (r ∷ʳ sp) = let (_ , eq , pwl , pwr) = break sp in - -, ≡.refl ∷ʳ eq , pwl , r ∷ pwr + -, refl ∷ʳ eq , pwl , r ∷ pwr -- map diff --git a/src/Data/List/Relation/Unary/Any/Properties.agda b/src/Data/List/Relation/Unary/Any/Properties.agda index 30dbc6850b..a6eaa9d480 100644 --- a/src/Data/List/Relation/Unary/Any/Properties.agda +++ b/src/Data/List/Relation/Unary/Any/Properties.agda @@ -41,8 +41,8 @@ open import Function.Related.Propositional as Related using (Kind; Related) open import Level using (Level) open import Relation.Binary.Core using (Rel; REL) open import Relation.Binary.Definitions as B -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core + using (_≡_; refl; sym; trans; cong; cong₂; subst) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) open import Relation.Unary as U @@ -100,13 +100,13 @@ Any-cong {P = P} {Q = Q} {xs = xs} {ys} P↔Q xs≈ys = map-id : ∀ (f : P ⋐ P) → (∀ {x} (p : P x) → f p ≡ p) → (p : Any P xs) → Any.map f p ≡ p -map-id f hyp (here p) = ≡.cong here (hyp p) -map-id f hyp (there p) = ≡.cong there $ map-id f hyp p +map-id f hyp (here p) = cong here (hyp p) +map-id f hyp (there p) = cong there $ map-id f hyp p map-∘ : ∀ (f : Q ⋐ R) (g : P ⋐ Q) (p : Any P xs) → Any.map (f ∘ g) p ≡ Any.map f (Any.map g p) map-∘ f g (here p) = refl -map-∘ f g (there p) = ≡.cong there $ map-∘ f g p +map-∘ f g (there p) = cong there $ map-∘ f g p ------------------------------------------------------------------------ -- Any.lookup @@ -133,16 +133,16 @@ swap-there : ∀ {P : A → B → Set ℓ} → (any : Any (λ x → Any (P x) ys) xs) → swap (Any.map (there {x = x}) any) ≡ there (swap any) swap-there (here pys) = refl -swap-there (there pxys) = ≡.cong (Any.map there) (swap-there pxys) +swap-there (there pxys) = cong (Any.map there) (swap-there pxys) swap-invol : ∀ {P : A → B → Set ℓ} → (any : Any (λ x → Any (P x) ys) xs) → swap (swap any) ≡ any swap-invol (here (here px)) = refl swap-invol (here (there pys)) = - ≡.cong (Any.map there) (swap-invol (here pys)) + cong (Any.map there) (swap-invol (here pys)) swap-invol (there pxys) = - ≡.trans (swap-there (swap pxys)) (≡.cong there (swap-invol pxys)) + trans (swap-there (swap pxys)) (cong there (swap-invol pxys)) swap↔ : ∀ {P : A → B → Set ℓ} → Any (λ x → Any (P x) ys) xs ↔ Any (λ y → Any (flip P y) xs) ys @@ -228,31 +228,31 @@ Any-×⁻ pq with Product.map₂ (Product.map₂ find) (find pq) Any-×⁻ (Any-×⁺ (p , q)) - ≡⟨⟩ + ≡⟨⟩ (let (x , x∈xs , pq) = find (Any-×⁺ (p , q)) (y , y∈ys , p , q) = find pq in lose x∈xs p , lose y∈ys q) - ≡⟨ ≡.cong (λ • → let (x , x∈xs , pq) = • - (y , y∈ys , p , q) = find pq - in lose x∈xs p , lose y∈ys q) - (find∘map p (λ p → Any.map (p ,_) q)) ⟩ + ≡⟨ cong (λ • → let (x , x∈xs , pq) = • + (y , y∈ys , p , q) = find pq + in lose x∈xs p , lose y∈ys q) + (find∘map p (λ p → Any.map (p ,_) q)) ⟩ (let (x , x∈xs , p) = find p (y , y∈ys , p , q) = find (Any.map (p ,_) q) in lose x∈xs p , lose y∈ys q) - ≡⟨ ≡.cong (λ • → let (x , x∈xs , p) = find p - (y , y∈ys , p , q) = • - in lose x∈xs p , lose y∈ys q) - (find∘map q (proj₂ (proj₂ (find p)) ,_)) ⟩ + ≡⟨ cong (λ • → let (x , x∈xs , p) = find p + (y , y∈ys , p , q) = • + in lose x∈xs p , lose y∈ys q) + (find∘map q (proj₂ (proj₂ (find p)) ,_)) ⟩ (let (x , x∈xs , p) = find p (y , y∈ys , q) = find q in lose x∈xs p , lose y∈ys q) - ≡⟨ ≡.cong₂ _,_ (lose∘find p) (lose∘find q) ⟩ + ≡⟨ cong₂ _,_ (lose∘find p) (lose∘find q) ⟩ (p , q) ∎ @@ -263,16 +263,16 @@ Any-×⁻ pq with Product.map₂ (Product.map₂ find) (find pq) with find pq′ | (λ (f : (proj₁ (find pq′) ≡_) ⋐ _) → map∘find pq′ {f}) ... | (y , y∈ys , p , q) | lem₂ - rewrite ≡.sym $ map-∘ {R = λ x → Any (λ y → P x × Q y) ys} - (λ p → Any.map (λ q → p , q) (lose y∈ys q)) - (λ y → ≡.subst P y p) - x∈xs + rewrite sym $ map-∘ {R = λ x → Any (λ y → P x × Q y) ys} + (λ p → Any.map (λ q → p , q) (lose y∈ys q)) + (λ y → subst P y p) + x∈xs = lem₁ _ helper where helper : Any.map (λ q → p , q) (lose y∈ys q) ≡ pq′ - helper rewrite ≡.sym $ map-∘ (λ q → p , q) - (λ y → ≡.subst Q y q) - y∈ys + helper rewrite sym $ map-∘ (λ q → p , q) + (λ y → subst Q y q) + y∈ys = lem₂ _ refl ------------------------------------------------------------------------ @@ -317,12 +317,12 @@ module _ {f : A → B} where map⁺∘map⁻ : (p : Any P (List.map f xs)) → map⁺ (map⁻ p) ≡ p map⁺∘map⁻ {xs = x ∷ xs} (here p) = refl - map⁺∘map⁻ {xs = x ∷ xs} (there p) = ≡.cong there (map⁺∘map⁻ p) + map⁺∘map⁻ {xs = x ∷ xs} (there p) = cong there (map⁺∘map⁻ p) map⁻∘map⁺ : ∀ (P : Pred B p) → (p : Any (P ∘ f) xs) → map⁻ {P = P} (map⁺ p) ≡ p map⁻∘map⁺ P (here p) = refl - map⁻∘map⁺ P (there p) = ≡.cong there (map⁻∘map⁺ P p) + map⁻∘map⁺ P (there p) = cong there (map⁻∘map⁺ P p) map↔ : Any (P ∘ f) xs ↔ Any P (List.map f xs) map↔ = mk↔ₛ′ map⁺ map⁻ map⁺∘map⁻ (map⁻∘map⁺ _) @@ -363,8 +363,8 @@ module _ {P : A → Set p} where ++⁺∘++⁻ [] p = refl ++⁺∘++⁻ (x ∷ xs) (here p) = refl ++⁺∘++⁻ (x ∷ xs) (there p) with ++⁻ xs p | ++⁺∘++⁻ xs p - ... | inj₁ p′ | ih = ≡.cong there ih - ... | inj₂ p′ | ih = ≡.cong there ih + ... | inj₁ p′ | ih = cong there ih + ... | inj₂ p′ | ih = cong there ih ++⁻∘++⁺ : ∀ xs {ys} (p : Any P xs ⊎ Any P ys) → ++⁻ xs ([ ++⁺ˡ , ++⁺ʳ xs ]′ p) ≡ p @@ -441,7 +441,7 @@ module _ {P : A → Set p} where concat⁻∘concat⁺ (here p) = concat⁻∘++⁺ˡ _ p concat⁻∘concat⁺ (there {x = xs} {xs = xss} p) rewrite concat⁻∘++⁺ʳ xs xss (concat⁺ p) = - ≡.cong there $ concat⁻∘concat⁺ p + cong there $ concat⁻∘concat⁺ p concat↔ : ∀ {xss} → Any (Any P) xss ↔ Any P (concat xss) concat↔ {xss} = mk↔ₛ′ concat⁺ (concat⁻ xss) (concat⁺∘concat⁻ xss) concat⁻∘concat⁺ @@ -496,7 +496,7 @@ applyUpTo⁻ f {suc n} (there p) with applyUpTo⁻ (f ∘ suc) p applyDownFrom⁺ : ∀ f {i n} → P (f i) → i < n → Any P (applyDownFrom f n) applyDownFrom⁺ f {i} {suc n} p (s≤s i≤n) with i ≟ n -... | yes ≡.refl = here p +... | yes refl = here p ... | no i≢n = there (applyDownFrom⁺ f p (≤∧≢⇒< i≤n i≢n)) applyDownFrom⁻ : ∀ f {n} → Any P (applyDownFrom f n) → @@ -611,7 +611,7 @@ module _ {P : B → Set p} where mapWith∈⁺ f (mapWith∈⁻ xs f p) ≡ p to∘from (y ∷ xs) f (here p) = refl to∘from (y ∷ xs) f (there p) = - ≡.cong there $ to∘from xs (f ∘ there) p + cong there $ to∘from xs (f ∘ there) p ------------------------------------------------------------------------ -- reverse @@ -686,7 +686,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where Any (λ f → Any (P ∘ f) xs) fs ↔⟨ Any-cong (λ _ → Any-cong (λ _ → pure↔) (_ ∎)) (_ ∎) ⟩ Any (λ f → Any (Any P ∘ pure ∘ f) xs) fs ↔⟨ Any-cong (λ _ → >>=↔ ) (_ ∎) ⟩ Any (λ f → Any P (xs >>= pure ∘ f)) fs ↔⟨ >>=↔ ⟩ - Any P (fs >>= λ f → xs >>= λ x → pure (f x)) ≡⟨ ≡.cong (Any P) (Listₑ.Applicative.unfold-⊛ fs xs) ⟨ + Any P (fs >>= λ f → xs >>= λ x → pure (f x)) ≡⟨ cong (Any P) (Listₑ.Applicative.unfold-⊛ fs xs) ⟨ Any P (fs ⊛ xs) ∎ where open Related.EquationalReasoning @@ -706,7 +706,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where Any (λ x → Any (λ y → P (x , y)) ys) xs ↔⟨ pure↔ ⟩ Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (pure _,_) ↔⟨ ⊛↔ ⟩ Any (λ x, → Any (P ∘ x,) ys) (pure _,_ ⊛ xs) ↔⟨ ⊛↔ ⟩ - Any P (pure _,_ ⊛ xs ⊛ ys) ≡⟨ ≡.cong (Any P ∘′ (_⊛ ys)) (Listₑ.Applicative.unfold-<$> _,_ xs) ⟨ + Any P (pure _,_ ⊛ xs ⊛ ys) ≡⟨ cong (Any P ∘′ (_⊛ ys)) (Listₑ.Applicative.unfold-<$> _,_ xs) ⟨ Any P (xs ⊗ ys) ∎ where open Related.EquationalReasoning diff --git a/src/Data/Maybe/Relation/Unary/All.agda b/src/Data/Maybe/Relation/Unary/All.agda index 31e4d8fd83..f26a261f42 100644 --- a/src/Data/Maybe/Relation/Unary/All.agda +++ b/src/Data/Maybe/Relation/Unary/All.agda @@ -16,7 +16,7 @@ open import Data.Product.Base as Product using (_,_) open import Function.Base using (id; _∘′_) open import Function.Bundles using (_⇔_; mk⇔) open import Level -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; cong) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) open import Relation.Unary open import Relation.Nullary hiding (Irrelevant) import Relation.Nullary.Decidable as Dec @@ -114,7 +114,7 @@ module _ {a p} {A : Set a} {P : Pred A p} where irrelevant : Irrelevant P → Irrelevant (All P) irrelevant P-irrelevant (just p) (just q) = cong just (P-irrelevant p q) - irrelevant P-irrelevant nothing nothing = ≡.refl + irrelevant P-irrelevant nothing nothing = refl satisfiable : Satisfiable (All P) satisfiable = nothing , nothing diff --git a/src/Data/Maybe/Relation/Unary/Any.agda b/src/Data/Maybe/Relation/Unary/Any.agda index d8cf44376f..d379ce2998 100644 --- a/src/Data/Maybe/Relation/Unary/Any.agda +++ b/src/Data/Maybe/Relation/Unary/Any.agda @@ -13,7 +13,7 @@ open import Data.Product.Base as Product using (∃; _,_; -,_) open import Function.Base using (id) open import Function.Bundles using (_⇔_; mk⇔) open import Level -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; cong) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; cong) open import Relation.Unary open import Relation.Nullary hiding (Irrelevant) import Relation.Nullary.Decidable as Dec diff --git a/src/Data/Nat/LCM.agda b/src/Data/Nat/LCM.agda index 73b4f6e19b..fc8937577b 100644 --- a/src/Data/Nat/LCM.agda +++ b/src/Data/Nat/LCM.agda @@ -17,8 +17,8 @@ open import Data.Nat.Properties open import Data.Nat.GCD open import Data.Product.Base using (_×_; _,_; uncurry′; ∃) open import Data.Sum.Base using (_⊎_; inj₁; inj₂) -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_; refl; sym; trans; cong; cong₂) +open import Relation.Binary.PropositionalEquality.Core + using (_≡_; refl; sym; trans; cong; cong₂; subst) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) open import Relation.Nullary.Decidable using (False; fromWitnessFalse) @@ -32,8 +32,8 @@ private -- Definition lcm : ℕ → ℕ → ℕ -lcm zero n = zero -lcm m@(suc m-1) n = m * (n / gcd m n) +lcm zero n = zero +lcm m@(suc _) n = m * (n / gcd m n) where instance _ = gcd≢0ˡ {m} {n} ------------------------------------------------------------------------ @@ -62,7 +62,7 @@ n∣lcm[m,n] m@(suc m-1) n = begin lcm-least : ∀ {m n c} → m ∣ c → n ∣ c → lcm m n ∣ c lcm-least {zero} {n} {c} 0∣c _ = 0∣c -lcm-least {m@(suc _)} {n} {c} m∣c n∣c = ≡.subst (_∣ c) (sym (rearrange m n)) +lcm-least {m@(suc _)} {n} {c} m∣c n∣c = subst (_∣ c) (sym (rearrange m n)) (m∣n*o⇒m/n∣o gcd[m,n]∣m*n mn∣c*gcd) where instance _ = gcd≢0ˡ {m} {n} @@ -73,7 +73,7 @@ lcm-least {m@(suc _)} {n} {c} m∣c n∣c = ≡.subst (_∣ c) (sym (rearrange m mn∣c*gcd : m * n ∣ c * gcd m n mn∣c*gcd = begin - m * n ∣⟨ gcd-greatest (≡.subst (_∣ c * m) (*-comm n m) (*-monoˡ-∣ m n∣c)) (*-monoˡ-∣ n m∣c) ⟩ + m * n ∣⟨ gcd-greatest (subst (_∣ c * m) (*-comm n m) (*-monoˡ-∣ m n∣c)) (*-monoˡ-∣ n m∣c) ⟩ gcd (c * m) (c * n) ≡⟨ c*gcd[m,n]≡gcd[cm,cn] c m n ⟨ c * gcd m n ∎ diff --git a/src/Data/Product/Function/Dependent/Propositional/WithK.agda b/src/Data/Product/Function/Dependent/Propositional/WithK.agda index 920bb2c136..0e0f63268b 100644 --- a/src/Data/Product/Function/Dependent/Propositional/WithK.agda +++ b/src/Data/Product/Function/Dependent/Propositional/WithK.agda @@ -20,7 +20,6 @@ open import Level using (Level) open import Function open import Function.Properties.Injection open import Function.Properties.Inverse as Inverse -open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl) private variable diff --git a/src/Data/Product/Relation/Binary/Pointwise/Dependent.agda b/src/Data/Product/Relation/Binary/Pointwise/Dependent.agda index 96b2cf3ba1..7050e8bd3a 100644 --- a/src/Data/Product/Relation/Binary/Pointwise/Dependent.agda +++ b/src/Data/Product/Relation/Binary/Pointwise/Dependent.agda @@ -17,7 +17,7 @@ open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions as B open import Relation.Binary.Indexed.Heterogeneous as I using (IREL; IRel; IndexedSetoid; IsIndexedEquivalence) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_) ------------------------------------------------------------------------ -- Pointwise lifting diff --git a/src/Data/Sum/Relation/Binary/LeftOrder.agda b/src/Data/Sum/Relation/Binary/LeftOrder.agda index 368066e6db..d29a6a96ed 100644 --- a/src/Data/Sum/Relation/Binary/LeftOrder.agda +++ b/src/Data/Sum/Relation/Binary/LeftOrder.agda @@ -24,7 +24,7 @@ open import Relation.Binary.Structures using (IsPreorder; IsPartialOrder; IsStrictPartialOrder; IsTotalOrder; IsDecTotalOrder; IsStrictTotalOrder) open import Relation.Binary.Definitions using (Reflexive; Transitive; Asymmetric; Total; Decidable; Irreflexive; Antisymmetric; Trichotomous; _Respectsʳ_; _Respectsˡ_; _Respects₂_; tri<; tri>; tri≈) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_) ------------------------------------------------------------------------ -- Definition diff --git a/src/Data/Vec/Functional/Relation/Binary/Equality/Setoid.agda b/src/Data/Vec/Functional/Relation/Binary/Equality/Setoid.agda index b14a2eaee2..1102985013 100644 --- a/src/Data/Vec/Functional/Relation/Binary/Equality/Setoid.agda +++ b/src/Data/Vec/Functional/Relation/Binary/Equality/Setoid.agda @@ -7,7 +7,7 @@ {-# OPTIONS --cubical-compatible --safe #-} open import Data.Nat.Base using (ℕ) -open import Data.Vec.Functional as VF hiding (map) +open import Data.Vec.Functional hiding (map) open import Data.Vec.Functional.Relation.Binary.Pointwise using (Pointwise) import Data.Vec.Functional.Relation.Binary.Pointwise.Properties as PW diff --git a/src/Data/Vec/Properties/WithK.agda b/src/Data/Vec/Properties/WithK.agda index bcd914ad51..f587d16b64 100644 --- a/src/Data/Vec/Properties/WithK.agda +++ b/src/Data/Vec/Properties/WithK.agda @@ -12,7 +12,7 @@ module Data.Vec.Properties.WithK where open import Data.Nat.Base open import Data.Nat.Properties using (+-assoc) open import Data.Vec.Base -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) open import Relation.Binary.HeterogeneousEquality as ≅ using (_≅_; refl) ------------------------------------------------------------------------ @@ -24,7 +24,7 @@ module _ {a} {A : Set a} where (p q : xs [ i ]= x) → p ≡ q []=-irrelevant here here = refl []=-irrelevant (there xs[i]=x) (there xs[i]=x′) = - ≡.cong there ([]=-irrelevant xs[i]=x xs[i]=x′) + cong there ([]=-irrelevant xs[i]=x xs[i]=x′) ------------------------------------------------------------------------ -- _++_ diff --git a/src/Data/Vec/Recursive/Properties.agda b/src/Data/Vec/Recursive/Properties.agda index df34cab770..3f25aa19d1 100644 --- a/src/Data/Vec/Recursive/Properties.agda +++ b/src/Data/Vec/Recursive/Properties.agda @@ -14,7 +14,8 @@ open import Data.Product.Base open import Data.Vec.Recursive open import Data.Vec.Base using (Vec; _∷_) open import Function.Bundles using (_↔_; mk↔ₛ′) -open import Relation.Binary.PropositionalEquality.Core as ≡ +open import Relation.Binary.PropositionalEquality.Core + using (_≡_; refl; cong; cong₂) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) open ≡-Reasoning @@ -28,31 +29,31 @@ private -- Basic proofs cons-head-tail-identity : ∀ n (as : A ^ suc n) → cons n (head n as) (tail n as) ≡ as -cons-head-tail-identity 0 as = ≡.refl -cons-head-tail-identity (suc n) as = ≡.refl +cons-head-tail-identity 0 as = refl +cons-head-tail-identity (suc n) as = refl head-cons-identity : ∀ n a (as : A ^ n) → head n (cons n a as) ≡ a -head-cons-identity 0 a as = ≡.refl -head-cons-identity (suc n) a as = ≡.refl +head-cons-identity 0 a as = refl +head-cons-identity (suc n) a as = refl tail-cons-identity : ∀ n a (as : A ^ n) → tail n (cons n a as) ≡ as -tail-cons-identity 0 a as = ≡.refl -tail-cons-identity (suc n) a as = ≡.refl +tail-cons-identity 0 a as = refl +tail-cons-identity (suc n) a as = refl append-cons : ∀ m n a (xs : A ^ m) ys → append (suc m) n (cons m a xs) ys ≡ cons (m + n) a (append m n xs ys) -append-cons 0 n a xs ys = ≡.refl -append-cons (suc m) n a xs ys = ≡.refl +append-cons 0 n a xs ys = refl +append-cons (suc m) n a xs ys = refl append-splitAt-identity : ∀ m n (as : A ^ (m + n)) → uncurry (append m n) (splitAt m n as) ≡ as -append-splitAt-identity 0 n as = ≡.refl +append-splitAt-identity 0 n as = refl append-splitAt-identity (suc m) n as = begin let x = head (m + n) as in let (xs , ys) = splitAt m n (tail (m + n) as) in append (suc m) n (cons m (head (m + n) as) xs) ys ≡⟨ append-cons m n x xs ys ⟩ cons (m + n) x (append m n xs ys) - ≡⟨ ≡.cong (cons (m + n) x) (append-splitAt-identity m n (tail (m + n) as)) ⟩ + ≡⟨ cong (cons (m + n) x) (append-splitAt-identity m n (tail (m + n) as)) ⟩ cons (m + n) x (tail (m + n) as) ≡⟨ cons-head-tail-identity (m + n) as ⟩ as @@ -62,21 +63,21 @@ append-splitAt-identity (suc m) n as = begin -- Conversion to and from Vec fromVec∘toVec : ∀ n (xs : A ^ n) → fromVec (toVec n xs) ≡ xs -fromVec∘toVec 0 _ = ≡.refl +fromVec∘toVec 0 _ = refl fromVec∘toVec (suc n) xs = begin cons n (head n xs) (fromVec (toVec n (tail n xs))) - ≡⟨ ≡.cong (cons n (head n xs)) (fromVec∘toVec n (tail n xs)) ⟩ + ≡⟨ cong (cons n (head n xs)) (fromVec∘toVec n (tail n xs)) ⟩ cons n (head n xs) (tail n xs) ≡⟨ cons-head-tail-identity n xs ⟩ xs ∎ toVec∘fromVec : ∀ {n} (xs : Vec A n) → toVec n (fromVec xs) ≡ xs -toVec∘fromVec Vec.[] = ≡.refl +toVec∘fromVec Vec.[] = refl toVec∘fromVec {n = suc n} (x Vec.∷ xs) = begin head n (cons n x (fromVec xs)) Vec.∷ toVec n (tail n (cons n x (fromVec xs))) - ≡⟨ ≡.cong₂ (λ x xs → x Vec.∷ toVec n xs) hd-prf tl-prf ⟩ + ≡⟨ cong₂ (λ x xs → x Vec.∷ toVec n xs) hd-prf tl-prf ⟩ x Vec.∷ toVec n (fromVec xs) - ≡⟨ ≡.cong (x Vec.∷_) (toVec∘fromVec xs) ⟩ + ≡⟨ cong (x Vec.∷_) (toVec∘fromVec xs) ⟩ x Vec.∷ xs ∎ where diff --git a/src/Data/Vec/Relation/Binary/Lex/Strict.agda b/src/Data/Vec/Relation/Binary/Lex/Strict.agda index 6276cfdacb..52f529aec9 100644 --- a/src/Data/Vec/Relation/Binary/Lex/Strict.agda +++ b/src/Data/Vec/Relation/Binary/Lex/Strict.agda @@ -33,7 +33,7 @@ open import Relation.Binary.Definitions using (Irreflexive; _Respects₂_; _Respectsˡ_; _Respectsʳ_; Antisymmetric; Asymmetric; Symmetric; Trans; Decidable; Total; Trichotomous; Transitive; Irrelevant; tri≈; tri>; tri<) open import Relation.Binary.Consequences open import Relation.Binary.Construct.On as On using (wellFounded) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) open import Level using (Level; _⊔_) private @@ -103,8 +103,8 @@ module _ {_≈_ : Rel A ℓ₁} {_≺_ : Rel A ℓ₂} where <-cmp : ∀ {n} → Trichotomous _≋_ (_<_ {n}) <-cmp [] [] = tri≈ ¬[]<[] [] ¬[]<[] <-cmp (x ∷ xs) (y ∷ ys) with ≺-cmp x y - ... | tri< x≺y x≉y x⊁y = tri< (this x≺y ≡.refl) (x≉y ∘ head) (≰-this (x≉y ∘ ≈-sym) x⊁y) - ... | tri> x⊀y x≉y x≻y = tri> (≰-this x≉y x⊀y) (x≉y ∘ head) (this x≻y ≡.refl) + ... | tri< x≺y x≉y x⊁y = tri< (this x≺y refl) (x≉y ∘ head) (≰-this (x≉y ∘ ≈-sym) x⊁y) + ... | tri> x⊀y x≉y x≻y = tri> (≰-this x≉y x⊀y) (x≉y ∘ head) (this x≻y refl) ... | tri≈ x⊀y x≈y x⊁y with <-cmp xs ys ... | tri< xs _ _ x≻y = inj₂ (this x≻y ≡.refl) + ... | tri< x≺y _ _ = inj₁ (this x≺y refl) + ... | tri> _ _ x≻y = inj₂ (this x≻y refl) ... | tri≈ _ x≈y _ with ≤-total xs ys ... | inj₁ xsys = inj₂ (next (≈-sym x≈y) xs>ys) diff --git a/src/Data/Vec/Relation/Unary/All.agda b/src/Data/Vec/Relation/Unary/All.agda index 8371d36c78..1c20d9c478 100644 --- a/src/Data/Vec/Relation/Unary/All.agda +++ b/src/Data/Vec/Relation/Unary/All.agda @@ -21,7 +21,7 @@ open import Relation.Nullary.Decidable as Dec using (_×-dec_; yes; no) open import Relation.Unary hiding (_∈_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (_Respects_) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (subst) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong₂) private variable @@ -94,7 +94,7 @@ lookupWith : ∀[ P ⇒ Q ⇒ R ] → All P xs → (i : Any Q xs) → R (Any.loo lookupWith f pxs i = Product.uncurry f (lookupAny pxs i) lookup : All P xs → (∀ {x} → x ∈ₚ xs → P x) -lookup pxs = lookupWith (λ { px ≡.refl → px }) pxs +lookup pxs = lookupWith (λ { px refl → px }) pxs module _(S : Setoid a ℓ) {P : Pred (Setoid.Carrier S) p} where open Setoid S renaming (sym to sym₁) @@ -115,9 +115,9 @@ universal u [] = [] universal u (x ∷ xs) = u x ∷ universal u xs irrelevant : Irrelevant P → ∀ {n} → Irrelevant (All P {n}) -irrelevant irr [] [] = ≡.refl +irrelevant irr [] [] = refl irrelevant irr (px₁ ∷ pxs₁) (px₂ ∷ pxs₂) = - ≡.cong₂ _∷_ (irr px₁ px₂) (irrelevant irr pxs₁ pxs₂) + cong₂ _∷_ (irr px₁ px₂) (irrelevant irr pxs₁ pxs₂) satisfiable : Satisfiable P → ∀ {n} → Satisfiable (All P {n}) satisfiable (x , p) {zero} = [] , [] diff --git a/src/Data/Vec/Relation/Unary/Any/Properties.agda b/src/Data/Vec/Relation/Unary/Any/Properties.agda index e95ae39930..0fb5512677 100644 --- a/src/Data/Vec/Relation/Unary/Any/Properties.agda +++ b/src/Data/Vec/Relation/Unary/Any/Properties.agda @@ -30,7 +30,8 @@ open import Relation.Nullary.Negation using (¬_) open import Relation.Unary hiding (_∈_) open import Relation.Binary.Core using (Rel) open import Relation.Binary.Definitions using (_Respects_) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core + using (_≡_; refl; sym; trans; cong) private variable @@ -93,15 +94,15 @@ module _ {P : Pred A p} where map-id : ∀ {P : Pred A p} (f : P ⊆ P) {n xs} → (∀ {x} (p : P x) → f p ≡ p) → (p : Any P {n} xs) → Any.map f p ≡ p -map-id f hyp (here p) = ≡.cong here (hyp p) -map-id f hyp (there p) = ≡.cong there $ map-id f hyp p +map-id f hyp (here p) = cong here (hyp p) +map-id f hyp (there p) = cong there $ map-id f hyp p map-∘ : ∀ {P : Pred A p} {Q : A → Set q} {R : A → Set r} (f : Q ⊆ R) (g : P ⊆ Q) {n xs} (p : Any P {n} xs) → Any.map (f ∘ g) p ≡ Any.map f (Any.map g p) map-∘ f g (here p) = refl -map-∘ f g (there p) = ≡.cong there $ map-∘ f g p +map-∘ f g (there p) = cong there $ map-∘ f g p ------------------------------------------------------------------------ -- Swapping @@ -117,7 +118,7 @@ module _ {P : A → B → Set ℓ} where swap-there : ∀ {n m x xs ys} → (any : Any (λ x → Any (P x) {n} ys) {m} xs) → swap (Any.map (there {x = x}) any) ≡ there (swap any) swap-there (here pys) = refl - swap-there (there pxys) = ≡.cong (Any.map there) (swap-there pxys) + swap-there (there pxys) = cong (Any.map there) (swap-there pxys) module _ {P : A → B → Set ℓ} where @@ -125,9 +126,9 @@ module _ {P : A → B → Set ℓ} where (any : Any (λ x → Any (P x) ys) xs) → swap (swap any) ≡ any swap-invol (here (here _)) = refl - swap-invol (here (there pys)) = ≡.cong (Any.map there) (swap-invol (here pys)) - swap-invol (there pxys) = ≡.trans (swap-there (swap pxys)) - $ ≡.cong there (swap-invol pxys) + swap-invol (here (there pys)) = cong (Any.map there) (swap-invol (here pys)) + swap-invol (there pxys) = trans (swap-there (swap pxys)) + $ cong there (swap-invol pxys) module _ {P : A → B → Set ℓ} where @@ -237,12 +238,12 @@ module _ {f : A → B} where map⁺∘map⁻ : ∀ {P : Pred B p} {n} {xs : Vec A n} → (p : Any P (map f xs)) → map⁺ (map⁻ p) ≡ p map⁺∘map⁻ {xs = x ∷ xs} (here p) = refl - map⁺∘map⁻ {xs = x ∷ xs} (there p) = ≡.cong there (map⁺∘map⁻ p) + map⁺∘map⁻ {xs = x ∷ xs} (there p) = cong there (map⁺∘map⁻ p) map⁻∘map⁺ : ∀ (P : Pred B p) {n} {xs : Vec A n} → (p : Any (P ∘ f) xs) → map⁻ {P = P} (map⁺ p) ≡ p map⁻∘map⁺ P (here p) = refl - map⁻∘map⁺ P (there p) = ≡.cong there (map⁻∘map⁺ P p) + map⁻∘map⁺ P (there p) = cong there (map⁻∘map⁺ P p) map↔ : ∀ {P : Pred B p} {n} {xs : Vec A n} → Any (P ∘ f) xs ↔ Any P (map f xs) @@ -271,8 +272,8 @@ module _ {P : Pred A p} where ++⁺∘++⁻ [] p = refl ++⁺∘++⁻ (x ∷ xs) (here p) = refl ++⁺∘++⁻ (x ∷ xs) (there p) with ++⁻ xs p | ++⁺∘++⁻ xs p - ++⁺∘++⁻ (x ∷ xs) (there p) | inj₁ p′ | ih = ≡.cong there ih - ++⁺∘++⁻ (x ∷ xs) (there p) | inj₂ p′ | ih = ≡.cong there ih + ++⁺∘++⁻ (x ∷ xs) (there p) | inj₁ p′ | ih = cong there ih + ++⁺∘++⁻ (x ∷ xs) (there p) | inj₂ p′ | ih = cong there ih ++⁻∘++⁺ : ∀ {n m} (xs : Vec A n) {ys : Vec A m} (p : Any P xs ⊎ Any P ys) → ++⁻ xs ([ ++⁺ˡ , ++⁺ʳ xs ]′ p) ≡ p @@ -336,10 +337,10 @@ module _ {P : Pred A p} where concat⁺ (concat⁻ xss p) ≡ p concat⁺∘concat⁻ (xs ∷ xss) p with ++⁻ xs p in eq ... | inj₁ pxs - = ≡.trans (≡.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (≡.sym eq)) + = trans (cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (sym eq)) $ ++⁺∘++⁻ xs p ... | inj₂ pxss rewrite concat⁺∘concat⁻ xss pxss - = ≡.trans (≡.cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (≡.sym eq)) + = trans (cong [ ++⁺ˡ , ++⁺ʳ xs ]′ (sym eq)) $ ++⁺∘++⁻ xs p concat⁻∘concat⁺ : ∀ {n m} {xss : Vec (Vec A n) m} (p : Any (Any P) xss) → @@ -402,7 +403,7 @@ module _ {P : Pred B p} where mapWith∈⁺ f (mapWith∈⁻ xs f p) ≡ p to∘from (y ∷ xs) f (here p) = refl to∘from (y ∷ xs) f (there p) = - ≡.cong there $ to∘from xs (f ∘ there) p + cong there $ to∘from xs (f ∘ there) p ------------------------------------------------------------------------ -- _∷_ diff --git a/src/Effect/Applicative.agda b/src/Effect/Applicative.agda index c681a18133..ef7dd5cc2f 100644 --- a/src/Effect/Applicative.agda +++ b/src/Effect/Applicative.agda @@ -21,7 +21,7 @@ open import Effect.Functor as Fun using (RawFunctor) open import Function.Base using (const; flip; _∘′_) open import Level using (Level; suc; _⊔_) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_) private variable diff --git a/src/Effect/Applicative/Indexed.agda b/src/Effect/Applicative/Indexed.agda index 1aa639292a..d08a515086 100644 --- a/src/Effect/Applicative/Indexed.agda +++ b/src/Effect/Applicative/Indexed.agda @@ -15,7 +15,7 @@ open import Effect.Functor using (RawFunctor) open import Data.Product.Base using (_×_; _,_) open import Function.Base open import Level -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong₂) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) @@ -111,6 +111,6 @@ record Morphism {I : Set i} {F₁ F₂ : IFun I f} op (f A₁.<$> x) ≡ (f A₂.<$> op x) op-<$> f x = begin op (A₁._⊛_ (A₁.pure f) x) ≡⟨ op-⊛ _ _ ⟩ - A₂._⊛_ (op (A₁.pure f)) (op x) ≡⟨ ≡.cong₂ A₂._⊛_ (op-pure _) ≡.refl ⟩ + A₂._⊛_ (op (A₁.pure f)) (op x) ≡⟨ cong₂ A₂._⊛_ (op-pure _) refl ⟩ A₂._⊛_ (A₂.pure f) (op x) ∎ where open ≡-Reasoning diff --git a/src/Function/Endomorphism/Propositional.agda b/src/Function/Endomorphism/Propositional.agda index 1846061b17..c8dd685328 100644 --- a/src/Function/Endomorphism/Propositional.agda +++ b/src/Function/Endomorphism/Propositional.agda @@ -21,7 +21,7 @@ open import Data.Product.Base using (_,_) open import Function.Base using (id; _∘′_; _∋_) open import Function.Equality using (_⟨$⟩_) open import Relation.Binary.Core using (_Preserves_⟶_) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong; cong₂) import Relation.Binary.PropositionalEquality.Properties as ≡ import Function.Endomorphism.Setoid (≡.setoid A) as Setoid @@ -38,7 +38,7 @@ fromSetoidEndo = _⟨$⟩_ toSetoidEndo : Endo → Setoid.Endo toSetoidEndo f = record { _⟨$⟩_ = f - ; cong = ≡.cong f + ; cong = cong f } ------------------------------------------------------------------------ @@ -52,7 +52,7 @@ f ^ suc n = f ∘′ (f ^ n) ^-homo : ∀ f → Homomorphic₂ ℕ Endo _≡_ (f ^_) _+_ _∘′_ ^-homo f zero n = refl -^-homo f (suc m) n = ≡.cong (f ∘′_) (^-homo f m n) +^-homo f (suc m) n = cong (f ∘′_) (^-homo f m n) ------------------------------------------------------------------------ -- Structures @@ -60,7 +60,7 @@ f ^ suc n = f ∘′ (f ^ n) ∘-isMagma : IsMagma _≡_ (Op₂ Endo ∋ _∘′_) ∘-isMagma = record { isEquivalence = ≡.isEquivalence - ; ∙-cong = ≡.cong₂ _∘′_ + ; ∙-cong = cong₂ _∘′_ } ∘-magma : Magma _ _ @@ -89,7 +89,7 @@ f ^ suc n = f ∘′ (f ^ n) ^-isSemigroupMorphism : ∀ f → IsSemigroupMorphism +-semigroup ∘-semigroup (f ^_) ^-isSemigroupMorphism f = record - { ⟦⟧-cong = ≡.cong (f ^_) + { ⟦⟧-cong = cong (f ^_) ; ∙-homo = ^-homo f } diff --git a/src/Function/Related/TypeIsomorphisms.agda b/src/Function/Related/TypeIsomorphisms.agda index 48ed246bca..eafe0b50b6 100644 --- a/src/Function/Related/TypeIsomorphisms.agda +++ b/src/Function/Related/TypeIsomorphisms.agda @@ -28,7 +28,7 @@ open import Function.Bundles open import Function.Related.Propositional import Function.Construct.Identity as Identity open import Relation.Binary hiding (_⇔_) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) open import Relation.Nullary.Reflects using (invert) @@ -47,20 +47,20 @@ 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↔ₛ′ Product.assocʳ Product.assocˡ (λ _ → ≡.refl) (λ _ → ≡.refl) +Σ-assoc = mk↔ₛ′ Product.assocʳ Product.assocˡ (λ _ → refl) (λ _ → refl) -- × is commutative ×-comm : ∀ (A : Set a) (B : Set b) → (A × B) ↔ (B × A) -×-comm _ _ = mk↔ₛ′ Product.swap Product.swap (λ _ → ≡.refl) λ _ → ≡.refl +×-comm _ _ = mk↔ₛ′ Product.swap Product.swap (λ _ → refl) λ _ → refl -- × has ⊤ as its identity ×-identityˡ : ∀ ℓ → LeftIdentity {ℓ = ℓ} _↔_ ⊤ _×_ -×-identityˡ _ _ = mk↔ₛ′ proj₂ -,_ (λ _ → ≡.refl) (λ _ → ≡.refl) +×-identityˡ _ _ = mk↔ₛ′ proj₂ -,_ (λ _ → refl) (λ _ → refl) ×-identityʳ : ∀ ℓ → RightIdentity {ℓ = ℓ} _↔_ ⊤ _×_ -×-identityʳ _ _ = mk↔ₛ′ proj₁ (_, _) (λ _ → ≡.refl) (λ _ → ≡.refl) +×-identityʳ _ _ = mk↔ₛ′ proj₁ (_, _) (λ _ → refl) (λ _ → refl) ×-identity : ∀ ℓ → Identity _↔_ ⊤ _×_ ×-identity ℓ = ×-identityˡ ℓ , ×-identityʳ ℓ @@ -68,10 +68,10 @@ private -- × has ⊥ has its zero ×-zeroˡ : ∀ ℓ → LeftZero {ℓ = ℓ} _↔_ ⊥ _×_ -×-zeroˡ ℓ A = mk↔ₛ′ proj₁ < id , ⊥ₚ-elim > (λ _ → ≡.refl) (λ { () }) +×-zeroˡ ℓ A = mk↔ₛ′ proj₁ < id , ⊥ₚ-elim > (λ _ → refl) (λ { () }) ×-zeroʳ : ∀ ℓ → RightZero {ℓ = ℓ} _↔_ ⊥ _×_ -×-zeroʳ ℓ A = mk↔ₛ′ proj₂ < ⊥ₚ-elim , id > (λ _ → ≡.refl) (λ { () }) +×-zeroʳ ℓ A = mk↔ₛ′ proj₂ < ⊥ₚ-elim , id > (λ _ → refl) (λ { () }) ×-zero : ∀ ℓ → Zero _↔_ ⊥ _×_ ×-zero ℓ = ×-zeroˡ ℓ , ×-zeroʳ ℓ @@ -85,8 +85,8 @@ private ⊎-assoc ℓ _ _ _ = mk↔ₛ′ [ [ inj₁ , inj₂ ∘′ inj₁ ]′ , inj₂ ∘′ inj₂ ]′ [ inj₁ ∘′ inj₁ , [ inj₁ ∘′ inj₂ , inj₂ ]′ ]′ - [ (λ _ → ≡.refl) , [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] ] - [ [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] , (λ _ → ≡.refl) ] + [ (λ _ → refl) , [ (λ _ → refl) , (λ _ → refl) ] ] + [ [ (λ _ → refl) , (λ _ → refl) ] , (λ _ → refl) ] -- ⊎ is commutative @@ -96,12 +96,12 @@ private -- ⊎ has ⊥ as its identity ⊎-identityˡ : ∀ ℓ → LeftIdentity _↔_ (⊥ {ℓ}) _⊎_ -⊎-identityˡ _ _ = mk↔ₛ′ [ (λ ()) , id ]′ inj₂ (λ _ → ≡.refl) - [ (λ ()) , (λ _ → ≡.refl) ] +⊎-identityˡ _ _ = mk↔ₛ′ [ (λ ()) , id ]′ inj₂ (λ _ → refl) + [ (λ ()) , (λ _ → refl) ] ⊎-identityʳ : ∀ ℓ → RightIdentity _↔_ (⊥ {ℓ}) _⊎_ -⊎-identityʳ _ _ = mk↔ₛ′ [ id , (λ ()) ]′ inj₁ (λ _ → ≡.refl) - [ (λ _ → ≡.refl) , (λ ()) ] +⊎-identityʳ _ _ = mk↔ₛ′ [ id , (λ ()) ]′ inj₁ (λ _ → refl) + [ (λ _ → refl) , (λ ()) ] ⊎-identity : ∀ ℓ → Identity _↔_ ⊥ _⊎_ ⊎-identity ℓ = ⊎-identityˡ ℓ , ⊎-identityʳ ℓ @@ -115,15 +115,15 @@ private ×-distribˡ-⊎ ℓ _ _ _ = mk↔ₛ′ (uncurry λ x → [ inj₁ ∘′ (x ,_) , inj₂ ∘′ (x ,_) ]′) [ Product.map₂ inj₁ , Product.map₂ inj₂ ]′ - [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] - (uncurry λ _ → [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ]) + [ (λ _ → refl) , (λ _ → refl) ] + (uncurry λ _ → [ (λ _ → refl) , (λ _ → refl) ]) ×-distribʳ-⊎ : ∀ ℓ → _DistributesOverʳ_ {ℓ = ℓ} _↔_ _×_ _⊎_ ×-distribʳ-⊎ ℓ _ _ _ = mk↔ₛ′ (uncurry [ curry inj₁ , curry inj₂ ]′) [ Product.map₁ inj₁ , Product.map₁ inj₂ ]′ - [ (λ _ → ≡.refl) , (λ _ → ≡.refl) ] - (uncurry [ (λ _ _ → ≡.refl) , (λ _ _ → ≡.refl) ]) + [ (λ _ → refl) , (λ _ → refl) ] + (uncurry [ (λ _ _ → refl) , (λ _ _ → refl) ]) ×-distrib-⊎ : ∀ ℓ → _DistributesOver_ {ℓ = ℓ} _↔_ _×_ _⊎_ ×-distrib-⊎ ℓ = ×-distribˡ-⊎ ℓ , ×-distribʳ-⊎ ℓ @@ -244,11 +244,11 @@ private ΠΠ↔ΠΠ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) → ((x : A) (y : B) → P x y) ↔ ((y : B) (x : A) → P x y) -ΠΠ↔ΠΠ _ = mk↔ₛ′ flip flip (λ _ → ≡.refl) (λ _ → ≡.refl) +ΠΠ↔ΠΠ _ = mk↔ₛ′ flip flip (λ _ → refl) (λ _ → refl) ∃∃↔∃∃ : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) → (∃₂ λ x y → P x y) ↔ (∃₂ λ y x → P x y) -∃∃↔∃∃ P = mk↔ₛ′ to from (λ _ → ≡.refl) (λ _ → ≡.refl) +∃∃↔∃∃ P = mk↔ₛ′ to from (λ _ → refl) (λ _ → refl) where to : (∃₂ λ x y → P x y) → (∃₂ λ y x → P x y) to (x , y , Pxy) = (y , x , Pxy) @@ -261,7 +261,7 @@ private Π↔Π : ∀ {A : Set a} {B : A → Set b} → ((x : A) → B x) ↔ ({x : A} → B x) -Π↔Π = mk↔ₛ′ _$- λ- (λ _ → ≡.refl) (λ _ → ≡.refl) +Π↔Π = mk↔ₛ′ _$- λ- (λ _ → refl) (λ _ → refl) ------------------------------------------------------------------------ -- _→_ preserves the symmetric relations @@ -280,12 +280,12 @@ private (λ f → from C↔D ∘ f ∘ to A↔B) (λ f → ext₂ λ x → begin to C↔D (from C↔D (f (to A↔B (from A↔B x)))) ≡⟨ strictlyInverseˡ C↔D _ ⟩ - f (to A↔B (from A↔B x)) ≡⟨ ≡.cong f $ strictlyInverseˡ A↔B x ⟩ + f (to A↔B (from A↔B x)) ≡⟨ cong f $ strictlyInverseˡ A↔B x ⟩ f x ∎) (λ f → ext₁ λ x → begin - from C↔D (to C↔D (f (from A↔B (to A↔B x)))) ≡⟨ strictlyInverseʳ C↔D _ ⟩ - f (from A↔B (to A↔B x)) ≡⟨ ≡.cong f $ strictlyInverseʳ A↔B x ⟩ - f x ∎) + from C↔D (to C↔D (f (from A↔B (to A↔B x)))) ≡⟨ strictlyInverseʳ C↔D _ ⟩ + f (from A↔B (to A↔B x)) ≡⟨ cong f $ strictlyInverseʳ A↔B x ⟩ + f x ∎) where open Inverse; open ≡-Reasoning →-cong : Extensionality a c → Extensionality b d → @@ -303,7 +303,7 @@ private ¬-cong : Extensionality a 0ℓ → Extensionality b 0ℓ → ∀ {k} {A : Set a} {B : Set b} → A ∼[ ⌊ k ⌋ ] B → (¬ A) ∼[ ⌊ k ⌋ ] (¬ B) -¬-cong extA extB A≈B = →-cong extA extB A≈B (K-reflexive ≡.refl) +¬-cong extA extB A≈B = →-cong extA extB A≈B (K-reflexive refl) ------------------------------------------------------------------------ -- _⇔_ preserves _⇔_ @@ -330,6 +330,6 @@ Related-cong {A = A} {B = B} {C = C} {D = D} A≈B C≈D = mk⇔ True↔ : ∀ {p} {P : Set p} (dec : Dec P) → ((p₁ p₂ : P) → p₁ ≡ p₂) → True dec ↔ P True↔ ( true because [p]) irr = - mk↔ₛ′ (λ _ → invert [p]) (λ _ → _) (irr _) (λ _ → ≡.refl) + mk↔ₛ′ (λ _ → invert [p]) (λ _ → _) (irr _) (λ _ → refl) True↔ (false because ofⁿ ¬p) _ = mk↔ₛ′ (λ()) (invert (ofⁿ ¬p)) (⊥-elim ∘ ¬p) (λ ()) diff --git a/src/Relation/Binary/Construct/Add/Infimum/NonStrict.agda b/src/Relation/Binary/Construct/Add/Infimum/NonStrict.agda index a151fe5065..c3b186c62a 100644 --- a/src/Relation/Binary/Construct/Add/Infimum/NonStrict.agda +++ b/src/Relation/Binary/Construct/Add/Infimum/NonStrict.agda @@ -20,8 +20,7 @@ module Relation.Binary.Construct.Add.Infimum.NonStrict open import Level using (_⊔_) open import Data.Sum.Base as Sum -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) import Relation.Binary.PropositionalEquality.Properties as ≡ import Relation.Binary.Construct.Add.Infimum.Equality as Equality open import Relation.Nullary hiding (Irrelevant) @@ -60,8 +59,8 @@ data _≤₋_ : Rel (A ₋) (a ⊔ ℓ) where ≤₋-total ≤-total [ k ] [ l ] = Sum.map [_] [_] (≤-total k l) ≤₋-irrelevant : Irrelevant _≤_ → Irrelevant _≤₋_ -≤₋-irrelevant ≤-irr (⊥₋≤ k) (⊥₋≤ k) = ≡.refl -≤₋-irrelevant ≤-irr [ p ] [ q ] = ≡.cong _ (≤-irr p q) +≤₋-irrelevant ≤-irr (⊥₋≤ k) (⊥₋≤ k) = refl +≤₋-irrelevant ≤-irr [ p ] [ q ] = cong _ (≤-irr p q) ------------------------------------------------------------------------ -- Relational properties + propositional equality @@ -72,7 +71,7 @@ data _≤₋_ : Rel (A ₋) (a ⊔ ℓ) where ≤₋-antisym-≡ : Antisymmetric _≡_ _≤_ → Antisymmetric _≡_ _≤₋_ ≤₋-antisym-≡ antisym (⊥₋≤ _) (⊥₋≤ _) = refl -≤₋-antisym-≡ antisym [ p ] [ q ] = ≡.cong [_] (antisym p q) +≤₋-antisym-≡ antisym [ p ] [ q ] = cong [_] (antisym p q) ------------------------------------------------------------------------ -- Relational properties + setoid equality diff --git a/src/Relation/Binary/Construct/Add/Infimum/Strict.agda b/src/Relation/Binary/Construct/Add/Infimum/Strict.agda index b9d47771db..c1312d7b14 100644 --- a/src/Relation/Binary/Construct/Add/Infimum/Strict.agda +++ b/src/Relation/Binary/Construct/Add/Infimum/Strict.agda @@ -21,8 +21,7 @@ module Relation.Binary.Construct.Add.Infimum.Strict open import Level using (_⊔_) open import Data.Product.Base using (_,_; map) open import Function.Base -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong; subst) import Relation.Binary.PropositionalEquality.Properties as ≡ import Relation.Binary.Construct.Add.Infimum.Equality as Equality import Relation.Binary.Construct.Add.Infimum.NonStrict as NonStrict @@ -59,8 +58,8 @@ data _<₋_ : Rel (A ₋) (a ⊔ ℓ) where <₋-dec _; _Respectsˡ_; _Respectsʳ_; _Respects_; _Respects₂_) open import Relation.Binary.Construct.Closure.Reflexive -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) import Relation.Binary.PropositionalEquality.Properties as ≡ open import Relation.Nullary import Relation.Nullary.Decidable as Dec diff --git a/src/Relation/Binary/Indexed/Heterogeneous/Bundles.agda b/src/Relation/Binary/Indexed/Heterogeneous/Bundles.agda index 78811d8edd..e1d8e53b07 100644 --- a/src/Relation/Binary/Indexed/Heterogeneous/Bundles.agda +++ b/src/Relation/Binary/Indexed/Heterogeneous/Bundles.agda @@ -11,10 +11,7 @@ module Relation.Binary.Indexed.Heterogeneous.Bundles where -open import Function.Base open import Level using (suc; _⊔_) -open import Relation.Binary.Core using (_⇒_) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Binary.Indexed.Heterogeneous.Core open import Relation.Binary.Indexed.Heterogeneous.Structures diff --git a/src/Relation/Binary/Indexed/Heterogeneous/Core.agda b/src/Relation/Binary/Indexed/Heterogeneous/Core.agda index 1be335bbfb..1015ef5449 100644 --- a/src/Relation/Binary/Indexed/Heterogeneous/Core.agda +++ b/src/Relation/Binary/Indexed/Heterogeneous/Core.agda @@ -14,7 +14,6 @@ module Relation.Binary.Indexed.Heterogeneous.Core where open import Level import Relation.Binary.Core as B import Relation.Binary.Definitions as B -import Relation.Binary.PropositionalEquality.Core as ≡ ------------------------------------------------------------------------ -- Indexed binary relations diff --git a/src/Relation/Binary/Indexed/Heterogeneous/Definitions.agda b/src/Relation/Binary/Indexed/Heterogeneous/Definitions.agda index e14345b83c..0fb09c6dd6 100644 --- a/src/Relation/Binary/Indexed/Heterogeneous/Definitions.agda +++ b/src/Relation/Binary/Indexed/Heterogeneous/Definitions.agda @@ -12,9 +12,7 @@ module Relation.Binary.Indexed.Heterogeneous.Definitions where open import Level -import Relation.Binary.Core as B import Relation.Binary.Definitions as B -import Relation.Binary.PropositionalEquality.Core as ≡ open import Relation.Binary.Indexed.Heterogeneous.Core private diff --git a/src/Relation/Binary/Indexed/Homogeneous/Bundles.agda b/src/Relation/Binary/Indexed/Homogeneous/Bundles.agda index 4fccb7fca9..aa56c0dcdb 100644 --- a/src/Relation/Binary/Indexed/Homogeneous/Bundles.agda +++ b/src/Relation/Binary/Indexed/Homogeneous/Bundles.agda @@ -11,12 +11,9 @@ module Relation.Binary.Indexed.Homogeneous.Bundles where -open import Data.Product.Base using (_,_) -open import Function.Base using (_⟨_⟩_) -open import Level using (Level; _⊔_; suc) -open import Relation.Binary.Core using (_⇒_; Rel) +open import Level using (suc; _⊔_) +open import Relation.Binary.Core using (Rel) open import Relation.Binary.Bundles as B -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) open import Relation.Nullary.Negation using (¬_) open import Relation.Binary.Indexed.Homogeneous.Core open import Relation.Binary.Indexed.Homogeneous.Structures diff --git a/src/Tactic/Cong.agda b/src/Tactic/Cong.agda index bf8a743306..2c02ec6a4c 100644 --- a/src/Tactic/Cong.agda +++ b/src/Tactic/Cong.agda @@ -36,7 +36,7 @@ open import Data.Unit.Base using (⊤) open import Data.Word.Base as Word using (toℕ) open import Data.Product.Base using (_×_; map₁; _,_) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl; cong) +open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) -- 'Data.String.Properties' defines this via 'Dec', so let's use the -- builtin for maximum speed.