Skip to content

Commit 6027dda

Browse files
authored
Qualified import of PropositionalEquality etc. fixing #2280 (#2293)
* Qualified import of `PropositionalEquality` fixing #2280 * Qualified import of `PropositionalEquality.Core` fixing #2280 * Qualified import of `HeterogeneousEquality.Core` fixing #2280 * simplified imports; fixed `README` link * simplified imports
1 parent 711ac53 commit 6027dda

File tree

115 files changed

+792
-805
lines changed

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

115 files changed

+792
-805
lines changed

src/Algebra/Construct/LiftedChoice.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -17,7 +17,7 @@ open import Relation.Nullary using (¬_; yes; no)
1717
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_])
1818
open import Data.Product.Base using (_×_; _,_)
1919
open import Level using (Level; _⊔_)
20-
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
20+
open import Relation.Binary.PropositionalEquality.Core as using (_≡_)
2121
open import Relation.Unary using (Pred)
2222

2323
import Relation.Binary.Reasoning.Setoid as ≈-Reasoning
@@ -55,8 +55,8 @@ module _ {_≈_ : Rel B ℓ} {_∙_ : Op₂ B}
5555

5656
sel-≡ : Selective _≡_ _◦_
5757
sel-≡ x y with M.sel (f x) (f y)
58-
... | inj₁ _ = inj₁ P.refl
59-
... | inj₂ _ = inj₂ P.refl
58+
... | inj₁ _ = inj₁ .refl
59+
... | inj₂ _ = inj₂ .refl
6060

6161
distrib : x y ((f x) ∙ (f y)) ≈ f (x ◦ y)
6262
distrib x y with M.sel (f x) (f y)

src/Algebra/Lattice/Properties/BooleanAlgebra/Expression.agda

+8-8
Original file line numberDiff line numberDiff line change
@@ -26,7 +26,7 @@ open import Data.Vec.Properties using (lookup-map)
2626
open import Data.Vec.Relation.Binary.Pointwise.Extensional as PW
2727
using (Pointwise; ext)
2828
open import Function.Base using (_∘_; _$_; flip)
29-
open import Relation.Binary.PropositionalEquality as P using (_≗_)
29+
open import Relation.Binary.PropositionalEquality as using (_≗_)
3030
import Relation.Binary.Reflection as Reflection
3131

3232
-- Expressions made up of variables and the operations of a boolean
@@ -68,7 +68,7 @@ module Naturality
6868
(f : Applicative.Morphism A₁ A₂)
6969
where
7070

71-
open P.≡-Reasoning
71+
open .≡-Reasoning
7272
open Applicative.Morphism f
7373
open Semantics A₁ renaming (⟦_⟧ to ⟦_⟧₁)
7474
open Semantics A₂ renaming (⟦_⟧ to ⟦_⟧₂)
@@ -77,21 +77,21 @@ module Naturality
7777

7878
natural : {n} (e : Expr n) op ∘ ⟦ e ⟧₁ ≗ ⟦ e ⟧₂ ∘ Vec.map op
7979
natural (var x) ρ = begin
80-
op (Vec.lookup ρ x) ≡⟨ P.sym $ lookup-map x op ρ ⟩
80+
op (Vec.lookup ρ x) ≡⟨ .sym $ lookup-map x op ρ ⟩
8181
Vec.lookup (Vec.map op ρ) x ∎
8282
natural (e₁ or e₂) ρ = begin
8383
op (_∨_ <$>₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩
84-
op (_∨_ <$>₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-<$> _ _) P.refl ⟩
85-
_∨_ <$>₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ (λ e₁ e₂ _∨_ <$>₂ e₁ ⊛₂ e₂) (natural e₁ ρ) (natural e₂ ρ) ⟩
84+
op (_∨_ <$>₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ .cong₂ _⊛₂_ (op-<$> _ _) .refl ⟩
85+
_∨_ <$>₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ .cong₂ (λ e₁ e₂ _∨_ <$>₂ e₁ ⊛₂ e₂) (natural e₁ ρ) (natural e₂ ρ) ⟩
8686
_∨_ <$>₂ ⟦ e₁ ⟧₂ (Vec.map op ρ) ⊛₂ ⟦ e₂ ⟧₂ (Vec.map op ρ) ∎
8787
natural (e₁ and e₂) ρ = begin
8888
op (_∧_ <$>₁ ⟦ e₁ ⟧₁ ρ ⊛₁ ⟦ e₂ ⟧₁ ρ) ≡⟨ op-⊛ _ _ ⟩
89-
op (_∧_ <$>₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ _⊛₂_ (op-<$> _ _) P.refl ⟩
90-
_∧_ <$>₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ P.cong₂ (λ e₁ e₂ _∧_ <$>₂ e₁ ⊛₂ e₂) (natural e₁ ρ) (natural e₂ ρ) ⟩
89+
op (_∧_ <$>₁ ⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ .cong₂ _⊛₂_ (op-<$> _ _) .refl ⟩
90+
_∧_ <$>₂ op (⟦ e₁ ⟧₁ ρ) ⊛₂ op (⟦ e₂ ⟧₁ ρ) ≡⟨ .cong₂ (λ e₁ e₂ _∧_ <$>₂ e₁ ⊛₂ e₂) (natural e₁ ρ) (natural e₂ ρ) ⟩
9191
_∧_ <$>₂ ⟦ e₁ ⟧₂ (Vec.map op ρ) ⊛₂ ⟦ e₂ ⟧₂ (Vec.map op ρ) ∎
9292
natural (not e) ρ = begin
9393
op (¬_ <$>₁ ⟦ e ⟧₁ ρ) ≡⟨ op-<$> _ _ ⟩
94-
¬_ <$>₂ op (⟦ e ⟧₁ ρ) ≡⟨ P.cong (¬_ <$>₂_) (natural e ρ) ⟩
94+
¬_ <$>₂ op (⟦ e ⟧₁ ρ) ≡⟨ .cong (¬_ <$>₂_) (natural e ρ) ⟩
9595
¬_ <$>₂ ⟦ e ⟧₂ (Vec.map op ρ) ∎
9696
natural top ρ = begin
9797
op (pure₁ ⊤) ≡⟨ op-pure _ ⟩

src/Algebra/Operations/CommutativeMonoid.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,7 @@ open import Data.Fin.Base using (Fin; zero)
1212
open import Data.Nat.Base as ℕ using (ℕ; zero; suc)
1313
open import Function.Base using (_∘_)
1414
open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_)
15-
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
15+
open import Relation.Binary.PropositionalEquality.Core as using (_≡_)
1616

1717
module Algebra.Operations.CommutativeMonoid
1818
{s₁ s₂} (CM : CommutativeMonoid s₁ s₂)
@@ -58,7 +58,7 @@ suc n ×′ x = x + n ×′ x
5858
×-congʳ (suc n) x≈x′ = +-cong x≈x′ (×-congʳ n x≈x′)
5959

6060
×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_
61-
×-cong {u} P.refl x≈x′ = ×-congʳ u x≈x′
61+
×-cong {u} .refl x≈x′ = ×-congʳ u x≈x′
6262

6363
-- _×_ is homomorphic with respect to _ℕ+_/_+_.
6464

@@ -98,7 +98,7 @@ suc n ×′ x = x + n ×′ x
9898
-- _×′_ preserves equality.
9999

100100
×′-cong : _×′_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_
101-
×′-cong {n} {_} {x} {y} P.refl x≈y = begin
101+
×′-cong {n} {_} {x} {y} .refl x≈y = begin
102102
n ×′ x ≈⟨ sym (×≈×′ n x) ⟩
103103
n × x ≈⟨ ×-congʳ n x≈y ⟩
104104
n × y ≈⟨ ×≈×′ n y ⟩

src/Algebra/Properties/CommutativeMonoid/Sum.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ open import Data.Fin.Permutation as Perm using (Permutation; _⟨$⟩ˡ_; _⟨$
1414
open import Data.Fin.Patterns using (0F)
1515
open import Data.Vec.Functional
1616
open import Function.Base using (_∘_)
17-
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
17+
open import Relation.Binary.PropositionalEquality.Core as using (_≡_)
1818
open import Relation.Nullary.Negation using (contradiction)
1919

2020
module Algebra.Properties.CommutativeMonoid.Sum
@@ -90,9 +90,9 @@ sum-permute {zero} {suc n} f π = contradiction π (Perm.refute λ())
9090
sum-permute {suc m} {zero} f π = contradiction π (Perm.refute λ())
9191
sum-permute {suc m} {suc n} f π = begin
9292
sum f ≡⟨⟩
93-
f 0F + sum f/0 ≡⟨ P.cong (_+ sum f/0) (P.cong f (Perm.inverseʳ π)) ⟨
93+
f 0F + sum f/0 ≡⟨ .cong (_+ sum f/0) (.cong f (Perm.inverseʳ π)) ⟨
9494
πf π₀ + sum f/0 ≈⟨ +-congˡ (sum-permute f/0 (Perm.remove π₀ π)) ⟩
95-
πf π₀ + sum (rearrange (π/0 ⟨$⟩ʳ_) f/0) ≡⟨ P.cong (πf π₀ +_) (sum-cong-≗ (P.cong f ∘ Perm.punchIn-permute′ π 0F)) ⟨
95+
πf π₀ + sum (rearrange (π/0 ⟨$⟩ʳ_) f/0) ≡⟨ .cong (πf π₀ +_) (sum-cong-≗ (.cong f ∘ Perm.punchIn-permute′ π 0F)) ⟨
9696
πf π₀ + sum (removeAt πf π₀) ≈⟨ sym (sum-remove πf) ⟩
9797
sum πf ∎
9898
where

src/Algebra/Properties/Monoid/Mult.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@
99
open import Algebra.Bundles using (Monoid)
1010
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; NonZero)
1111
open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_)
12-
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
12+
open import Relation.Binary.PropositionalEquality.Core as using (_≡_)
1313

1414
module Algebra.Properties.Monoid.Mult {a ℓ} (M : Monoid a ℓ) where
1515

@@ -44,7 +44,7 @@ open import Algebra.Definitions.RawMonoid rawMonoid public
4444
×-congʳ (suc n) x≈x′ = +-cong x≈x′ (×-congʳ n x≈x′)
4545

4646
×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_
47-
×-cong {n} P.refl x≈x′ = ×-congʳ n x≈x′
47+
×-cong {n} .refl x≈x′ = ×-congʳ n x≈x′
4848

4949
×-congˡ : {x} (_× x) Preserves _≡_ ⟶ _≈_
5050
×-congˡ m≡n = ×-cong m≡n refl

src/Algebra/Properties/Monoid/Mult/TCOptimised.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,7 @@
1010
open import Algebra.Bundles using (Monoid)
1111
open import Data.Nat.Base as ℕ using (ℕ; zero; suc)
1212
open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_)
13-
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
13+
open import Relation.Binary.PropositionalEquality.Core as using (_≡_)
1414

1515
module Algebra.Properties.Monoid.Mult.TCOptimised
1616
{a ℓ} (M : Monoid a ℓ) where
@@ -75,7 +75,7 @@ open import Algebra.Definitions.RawMonoid rawMonoid public
7575
×-congʳ (suc n@(suc _)) x≈y = +-cong (×-congʳ n x≈y) x≈y
7676

7777
×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_
78-
×-cong {n} P.refl x≈y = ×-congʳ n x≈y
78+
×-cong {n} .refl x≈y = ×-congʳ n x≈y
7979

8080
×-assocˡ : x m n m × (n × x) ≈ (m ℕ.* n) × x
8181
×-assocˡ x m n = begin

src/Algebra/Properties/Monoid/Sum.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -13,7 +13,7 @@ open import Data.Fin.Base using (zero; suc)
1313
open import Data.Unit using (tt)
1414
open import Function.Base using (_∘_)
1515
open import Relation.Binary.Core using (_Preserves_⟶_)
16-
open import Relation.Binary.PropositionalEquality as P using (_≗_; _≡_)
16+
open import Relation.Binary.PropositionalEquality as using (_≗_; _≡_)
1717

1818
module Algebra.Properties.Monoid.Sum {a ℓ} (M : Monoid a ℓ) where
1919

@@ -61,8 +61,8 @@ sum-cong-≋ {zero} xs≋ys = refl
6161
sum-cong-≋ {suc n} xs≋ys = +-cong (xs≋ys zero) (sum-cong-≋ (xs≋ys ∘ suc))
6262

6363
sum-cong-≗ : {n} sum {n} Preserves _≗_ ⟶ _≡_
64-
sum-cong-≗ {zero} xs≗ys = P.refl
65-
sum-cong-≗ {suc n} xs≗ys = P.cong₂ _+_ (xs≗ys zero) (sum-cong-≗ (xs≗ys ∘ suc))
64+
sum-cong-≗ {zero} xs≗ys = .refl
65+
sum-cong-≗ {suc n} xs≗ys = .cong₂ _+_ (xs≗ys zero) (sum-cong-≗ (xs≗ys ∘ suc))
6666

6767
sum-replicate : n {x} sum (replicate n x) ≈ n × x
6868
sum-replicate zero = refl

src/Algebra/Properties/Semiring/Exp.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@
99
open import Algebra
1010
open import Data.Nat.Base as ℕ using (ℕ; zero; suc)
1111
open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_)
12-
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
12+
open import Relation.Binary.PropositionalEquality.Core as using (_≡_)
1313
import Data.Nat.Properties as ℕ
1414

1515
module Algebra.Properties.Semiring.Exp

src/Algebra/Solver/Monoid.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -22,7 +22,7 @@ open import Data.Vec.Base using (Vec; lookup)
2222
open import Function.Base using (_∘_; _$_)
2323
open import Relation.Binary.Definitions using (Decidable)
2424

25-
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
25+
open import Relation.Binary.PropositionalEquality.Core using (_≡_; cong)
2626
import Relation.Binary.Reflection
2727
open import Relation.Nullary
2828
import Relation.Nullary.Decidable as Dec
@@ -128,7 +128,7 @@ prove′ e₁ e₂ =
128128
lemma : normalise e₁ ≡ normalise e₂ ρ ⟦ e₁ ⟧ ρ ≈ ⟦ e₂ ⟧ ρ
129129
lemma eq ρ =
130130
R.prove ρ e₁ e₂ (begin
131-
⟦ normalise e₁ ⟧⇓ ρ ≡⟨ P.cong (λ e ⟦ e ⟧⇓ ρ) eq ⟩
131+
⟦ normalise e₁ ⟧⇓ ρ ≡⟨ cong (λ e ⟦ e ⟧⇓ ρ) eq ⟩
132132
⟦ normalise e₂ ⟧⇓ ρ ∎)
133133

134134
-- This procedure can be combined with from-just.

src/Algebra/Solver/Ring.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -41,7 +41,7 @@ open import Algebra.Properties.Semiring.Exp semiring
4141

4242
open import Relation.Nullary.Decidable using (yes; no)
4343
open import Relation.Binary.Reasoning.Setoid setoid
44-
import Relation.Binary.PropositionalEquality.Core as PropEq
44+
import Relation.Binary.PropositionalEquality.Core as
4545
import Relation.Binary.Reflection as Reflection
4646

4747
open import Data.Nat.Base using (ℕ; suc; zero)
@@ -534,7 +534,7 @@ correct (con c) ρ = correct-con c ρ
534534
correct (var i) ρ = correct-var i ρ
535535
correct (p :^ k) ρ = begin
536536
⟦ normalise p ^N k ⟧N ρ ≈⟨ ^N-homo (normalise p) k ρ ⟩
537-
⟦ p ⟧↓ ρ ^ k ≈⟨ correct p ρ ⟨ ^-cong ⟩ PropEq.refl {x = k} ⟩
537+
⟦ p ⟧↓ ρ ^ k ≈⟨ correct p ρ ⟨ ^-cong ⟩ .refl {x = k} ⟩
538538
⟦ p ⟧ ρ ^ k ∎
539539
correct (:- p) ρ = begin
540540
⟦ -N normalise p ⟧N ρ ≈⟨ -N‿-homo (normalise p) ρ ⟩

src/Algebra/Solver/Ring/NaturalCoefficients/Default.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -20,13 +20,13 @@ import Algebra.Properties.Semiring.Mult as SemiringMultiplication
2020
open import Data.Maybe.Base using (Maybe; map)
2121
open import Data.Nat using (_≟_)
2222
open import Relation.Binary.Consequences using (dec⇒weaklyDec)
23-
import Relation.Binary.PropositionalEquality.Core as P
23+
import Relation.Binary.PropositionalEquality.Core as
2424

2525
open CommutativeSemiring R
2626
open SemiringMultiplication semiring
2727

2828
private
2929
dec : m n Maybe (m × 1# ≈ n × 1#)
30-
dec m n = map (λ { P.refl refl }) (dec⇒weaklyDec _≟_ m n)
30+
dec m n = map (λ { .refl refl }) (dec⇒weaklyDec _≟_ m n)
3131

3232
open import Algebra.Solver.Ring.NaturalCoefficients R dec public

0 commit comments

Comments
 (0)