Skip to content

Commit ae348fd

Browse files
Rename symmetric reasoning combinators. Minimal set of changes (#2160)
1 parent 2b2a130 commit ae348fd

File tree

77 files changed

+716
-572
lines changed

Some content is hidden

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

77 files changed

+716
-572
lines changed

CHANGELOG.md

+48-15
Original file line numberDiff line numberDiff line change
@@ -849,6 +849,7 @@ Non-backwards compatible changes
849849
IO.Effectful
850850
IO.Instances
851851
```
852+
852853
### (Issue #2096) Introduction of flipped relation symbol for `Relation.Binary.Bundles.Preorder`
853854

854855
* Previously, the relation symbol `_∼_` was (notationally) symmetric, so that its
@@ -989,6 +990,53 @@ Non-backwards compatible changes
989990
Relation.Binary.Reasoning.PartialOrder
990991
```
991992

993+
### More modular design of equational reasoning.
994+
995+
* Have introduced a new module `Relation.Binary.Reasoning.Syntax` which exports
996+
a range of modules containing pre-existing reasoning combinator syntax.
997+
998+
* This makes it possible to add new or rename existing reasoning combinators to a
999+
pre-existing `Reasoning` module in just a couple of lines
1000+
(e.g. see `∣-Reasoning` in `Data.Nat.Divisibility`)
1001+
1002+
* One pre-requisite for that is that `≡-Reasoning` has been moved from
1003+
`Relation.Binary.PropositionalEquality.Core` (which shouldn't be
1004+
imported anyway as it's a `Core` module) to
1005+
`Relation.Binary.PropositionalEquality.Properties`.
1006+
It is still exported by `Relation.Binary.PropositionalEquality`.
1007+
1008+
### Renaming of symmetric combinators for equational reasoning
1009+
1010+
* We've had various complaints about the symmetric version of reasoning combinators
1011+
that use the syntax `_R˘⟨_⟩_` for some relation `R`, (e.g. `_≡˘⟨_⟩_` and `_≃˘⟨_⟩_`)
1012+
introduced in `v1.0`. In particular:
1013+
1. The symbol `˘` is hard to type.
1014+
2. The symbol `˘` doesn't convey the direction of the equality
1015+
3. The left brackets aren't vertically aligned with the left brackets of the non-symmetric version.
1016+
1017+
* To address these problems we have renamed all the symmetric versions of the
1018+
combinators from `_R˘⟨_⟩_` to `_R⟨_⟨_` (the direction of the last bracket is flipped
1019+
to indicate the quality goes from right to left).
1020+
1021+
* The old combinators still exist but have been deprecated. However due to
1022+
[Agda issue #5617](https://github.com/agda/agda/issues/5617), the deprecation warnings
1023+
don't fire correctly. We will not remove the old combinators before the above issue is
1024+
addressed. However, we still encourage migrating to the new combinators!
1025+
1026+
* On a Linux-based system, the following command was used to globally migrate all uses of the
1027+
old combinators to the new ones in the standard library itself.
1028+
It *may* be useful when trying to migrate your own code:
1029+
```bash
1030+
find . -type f -name "*.agda" -print0 | xargs -0 sed -i -e 's/˘⟨\(.*\)⟩/⟨\1⟨/g'
1031+
```
1032+
USUAL CAVEATS: It has not been widely tested and the standard library developers are not
1033+
responsible for the results of this command. It is strongly recommended you back up your
1034+
work before you attempt to run it.
1035+
1036+
* NOTE: this refactoring may require some extra bracketing around the operator `_⟨_⟩_` from
1037+
`Function.Base` if the `_⟨_⟩_` operator is used within the reasoning proof. The symptom
1038+
for this is a `Don't know how to parse` error.
1039+
9921040
### Other
9931041

9941042
* In accordance with changes to the flags in Agda 2.6.3, all modules that previously used
@@ -1228,21 +1276,6 @@ Major improvements
12281276
* `RawLattice` has been moved from `Algebra.Lattice.Bundles` to this new module.
12291277

12301278
* In `Relation.Binary.Reasoning.Base.Triple`, added a new parameter `<-irrefl : Irreflexive _≈_ _<_`
1231-
1232-
### More modular design of equational reasoning.
1233-
1234-
* Have introduced a new module `Relation.Binary.Reasoning.Syntax` which exports
1235-
a range of modules containing pre-existing reasoning combinator syntax.
1236-
1237-
* This makes it possible to add new or rename existing reasoning combinators to a
1238-
pre-existing `Reasoning` module in just a couple of lines
1239-
(e.g. see `∣-Reasoning` in `Data.Nat.Divisibility`)
1240-
1241-
* One pre-requisite for that is that `≡-Reasoning` has been moved from
1242-
`Relation.Binary.PropositionalEquality.Core` (which shouldn't be
1243-
imported anyway as it's a `Core` module) to
1244-
`Relation.Binary.PropositionalEquality.Properties`.
1245-
It is still exported by `Relation.Binary.PropositionalEquality`.
12461279

12471280
Deprecated modules
12481281
------------------

README/Data/Fin/Relation/Unary/Top.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -76,8 +76,8 @@ opposite-prop {suc n} i with view i
7676
... | ‵fromℕ rewrite toℕ-fromℕ n | n∸n≡0 n = refl
7777
... | ‵inject₁ j = begin
7878
suc (toℕ (opposite j)) ≡⟨ cong suc (opposite-prop j) ⟩
79-
suc (n ∸ suc (toℕ j)) ≡˘⟨ +-∸-assoc 1 (toℕ<n j)
80-
n ∸ toℕ j ≡˘⟨ cong (n ∸_) (toℕ-inject₁ j)
79+
suc (n ∸ suc (toℕ j)) ≡⟨ +-∸-assoc 1 (toℕ<n j)
80+
n ∸ toℕ j ≡⟨ cong (n ∸_) (toℕ-inject₁ j)
8181
n ∸ toℕ (inject₁ j) ∎ where open ≡-Reasoning
8282

8383
------------------------------------------------------------------------

README/Data/Vec/Relation/Binary/Equality/Cast.agda

+12-12
Original file line numberDiff line numberDiff line change
@@ -86,7 +86,7 @@ example1a-fromList-∷ʳ : ∀ (x : A) xs .(eq : L.length (xs L.∷ʳ x) ≡ suc
8686
cast eq (fromList (xs L.∷ʳ x)) ≡ fromList xs ∷ʳ x
8787
example1a-fromList-∷ʳ x xs eq = begin
8888
cast eq (fromList (xs L.∷ʳ x)) ≡⟨⟩
89-
cast eq (fromList (xs L.++ L.[ x ])) ≡˘⟨ cast-trans eq₁ eq₂ (fromList (xs L.++ L.[ x ]))
89+
cast eq (fromList (xs L.++ L.[ x ])) ≡⟨ cast-trans eq₁ eq₂ (fromList (xs L.++ L.[ x ]))
9090
cast eq₂ (cast eq₁ (fromList (xs L.++ L.[ x ]))) ≡⟨ cong (cast eq₂) (fromList-++ xs) ⟩
9191
cast eq₂ (fromList xs ++ [ x ]) ≡⟨ ≈-sym (unfold-∷ʳ (sym eq₂) x (fromList xs)) ⟩
9292
fromList xs ∷ʳ x ∎
@@ -105,7 +105,7 @@ example1b-fromList-∷ʳ : ∀ (x : A) xs .(eq : L.length (xs L.∷ʳ x) ≡ suc
105105
example1b-fromList-∷ʳ x xs eq = begin
106106
fromList (xs L.∷ʳ x) ≈⟨⟩
107107
fromList (xs L.++ L.[ x ]) ≈⟨ fromList-++ xs ⟩
108-
fromList xs ++ [ x ] ≈˘⟨ unfold-∷ʳ (+-comm 1 (L.length xs)) x (fromList xs)
108+
fromList xs ++ [ x ] ≈⟨ unfold-∷ʳ (+-comm 1 (L.length xs)) x (fromList xs)
109109
fromList xs ∷ʳ x ∎
110110
where open CastReasoning
111111

@@ -120,10 +120,10 @@ example1b-fromList-∷ʳ x xs eq = begin
120120
-- ≈-trans : xs ≈[ m≡n ] ys → ys ≈[ n≡o ] zs → xs ≈[ trans m≡n n≡o ] zs
121121
-- Accordingly, `_≈[_]_` admits the standard set of equational reasoning
122122
-- combinators. Suppose `≈-eqn : xs ≈[ m≡n ] ys`,
123-
-- xs ≈⟨ ≈-eqn ⟩ -- `_≈⟨_⟩_` takes a `_≈[_]_` step, adjusting
124-
-- ys -- the index at the same time
123+
-- xs ≈⟨ ≈-eqn ⟩ -- `_≈⟨_⟩_` takes a `_≈[_]_` step, adjusting
124+
-- ys -- the index at the same time
125125
--
126-
-- ys ≈˘⟨ ≈-eqn -- `_≈˘⟨_⟩_` takes a symmetric `_≈[_]_` step
126+
-- ys ≈⟨ ≈-eqn -- `_≈⟨_⟨_` takes a symmetric `_≈[_]_` step
127127
-- xs
128128
example2a : .(eq : suc m + n ≡ m + suc n) (xs : Vec A m) a ys
129129
cast eq ((reverse xs ∷ʳ a) ++ ys) ≡ reverse xs ++ (a ∷ ys)
@@ -138,7 +138,7 @@ example2a eq xs a ys = begin
138138
-- xs ≂⟨ ≡-eqn ⟩ -- Takes a `_≡_` step; no change to the index
139139
-- ys
140140
--
141-
-- ys ≂˘⟨ ≡-eqn -- Takes a symmetric `_≡_` step
141+
-- ys ≂⟨ ≡-eqn -- Takes a symmetric `_≡_` step
142142
-- xs
143143
-- Equivalently, `≈-reflexive` injects `_≡_` into `_≈[_]_`. That is,
144144
-- `xs ≂⟨ ≡-eqn ⟩ ys` is the same as `xs ≈⟨ ≈-reflexive ≡-eqn ⟩ ys`.
@@ -149,7 +149,7 @@ example2b eq xs a ys = begin
149149
(a ∷ xs) ʳ++ ys ≂⟨ unfold-ʳ++ (a ∷ xs) ys ⟩ -- index: suc m + n
150150
reverse (a ∷ xs) ++ ys ≂⟨ cong (_++ ys) (reverse-∷ a xs) ⟩ -- index: suc m + n
151151
(reverse xs ∷ʳ a) ++ ys ≈⟨ ∷ʳ-++ eq a (reverse xs) ⟩ -- index: suc m + n
152-
reverse xs ++ (a ∷ ys) ≂˘⟨ unfold-ʳ++ xs (a ∷ ys) -- index: m + suc n
152+
reverse xs ++ (a ∷ ys) ≂⟨ unfold-ʳ++ xs (a ∷ ys) -- index: m + suc n
153153
xs ʳ++ (a ∷ ys) ∎ -- index: m + suc n
154154
where open CastReasoning
155155

@@ -201,11 +201,11 @@ example3b-fromList-++-++′ {xs = xs} {ys} {zs} eq = begin
201201
example4-cong² : .(eq : (m + 1) + n ≡ n + suc m) a (xs : Vec A m) ys
202202
cast eq (reverse ((xs ++ [ a ]) ++ ys)) ≡ ys ʳ++ reverse (xs ∷ʳ a)
203203
example4-cong² {m = m} {n} eq a xs ys = begin
204-
reverse ((xs ++ [ a ]) ++ ys) ≈˘⟨ ≈-cong reverse (cast-reverse (cong (_+ n) (+-comm 1 m)) ((xs ∷ʳ a) ++ ys))
204+
reverse ((xs ++ [ a ]) ++ ys) ≈⟨ ≈-cong reverse (cast-reverse (cong (_+ n) (+-comm 1 m)) ((xs ∷ʳ a) ++ ys))
205205
(≈-cong (_++ ys) (cast-++ˡ (+-comm 1 m) (xs ∷ʳ a))
206-
(unfold-∷ʳ _ a xs))
206+
(unfold-∷ʳ _ a xs))
207207
reverse ((xs ∷ʳ a) ++ ys) ≈⟨ reverse-++ (+-comm (suc m) n) (xs ∷ʳ a) ys ⟩
208-
reverse ys ++ reverse (xs ∷ʳ a) ≂˘⟨ unfold-ʳ++ ys (reverse (xs ∷ʳ a))
208+
reverse ys ++ reverse (xs ∷ʳ a) ≂⟨ unfold-ʳ++ ys (reverse (xs ∷ʳ a))
209209
ys ʳ++ reverse (xs ∷ʳ a) ∎
210210
where open CastReasoning
211211

@@ -237,7 +237,7 @@ example5-fromList-++-++′ {xs = xs} {ys} {zs} eq = begin
237237
-- and then switch to the reasoning system of `_≈[_]_`.
238238
example6a-reverse-∷ʳ : x (xs : Vec A n) reverse (xs ∷ʳ x) ≡ x ∷ reverse xs
239239
example6a-reverse-∷ʳ {n = n} x xs = begin-≡
240-
reverse (xs ∷ʳ x) ≡˘⟨ ≈-reflexive refl
240+
reverse (xs ∷ʳ x) ≡⟨ ≈-reflexive refl
241241
reverse (xs ∷ʳ x) ≈⟨ ≈-cong reverse (cast-reverse _ _) (unfold-∷ʳ (+-comm 1 n) x xs) ⟩
242242
reverse (xs ++ [ x ]) ≈⟨ reverse-++ (+-comm n 1) xs [ x ] ⟩
243243
x ∷ reverse xs ∎
@@ -249,6 +249,6 @@ example6b-reverse-∷ʳ-by-induction x (y ∷ xs) = begin
249249
reverse (y ∷ (xs ∷ʳ x)) ≡⟨ reverse-∷ y (xs ∷ʳ x) ⟩
250250
reverse (xs ∷ʳ x) ∷ʳ y ≡⟨ cong (_∷ʳ y) (example6b-reverse-∷ʳ-by-induction x xs) ⟩
251251
(x ∷ reverse xs) ∷ʳ y ≡⟨⟩
252-
x ∷ (reverse xs ∷ʳ y) ≡˘⟨ cong (x ∷_) (reverse-∷ y xs)
252+
x ∷ (reverse xs ∷ʳ y) ≡⟨ cong (x ∷_) (reverse-∷ y xs)
253253
x ∷ reverse (y ∷ xs) ∎
254254
where open ≡-Reasoning

README/Tactic/Cong.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -33,7 +33,7 @@ verbose-example m n eq =
3333
suc (suc m) + m
3434
≡⟨ cong (λ ϕ suc (suc (ϕ + ϕ))) eq ⟩
3535
suc (suc n) + n
36-
˘⟨ cong (λ ϕ suc (suc (n + ϕ))) (+-identityʳ n)
36+
≡⟨ cong (λ ϕ suc (suc (n + ϕ))) (+-identityʳ n)
3737
suc (suc n) + (n + 0)
3838
3939

@@ -51,7 +51,7 @@ succinct-example m n eq =
5151
suc (suc m) + m
5252
≡⟨ cong! eq ⟩
5353
suc (suc n) + n
54-
˘⟨ cong! (+-identityʳ n)
54+
≡⟨ cong! (+-identityʳ n)
5555
suc (suc n) + (n + 0)
5656
5757

@@ -124,7 +124,7 @@ module EquationalReasoningTests where
124124
suc (suc m) + m
125125
≡⟨ cong! eq ⟩
126126
suc (suc n) + n
127-
˘⟨ cong! (+-identityʳ n)
127+
≡⟨ cong! (+-identityʳ n)
128128
suc (suc n) + (n + 0)
129129
130130

src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda

+8-8
Original file line numberDiff line numberDiff line change
@@ -56,11 +56,11 @@ x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0)
5656
y⁻¹*x⁻¹*x*y≈1 = begin
5757
y⁻¹ * x⁻¹ * (x * y - 0#) ≈⟨ *-congˡ (x-0≈x (x * y)) ⟩
5858
y⁻¹ * x⁻¹ * (x * y) ≈⟨ *-assoc y⁻¹ x⁻¹ (x * y) ⟩
59-
y⁻¹ * (x⁻¹ * (x * y)) ≈˘⟨ *-congˡ (*-assoc x⁻¹ x y)
60-
y⁻¹ * ((x⁻¹ * x) * y) ≈˘⟨ *-congˡ (*-congʳ (*-congˡ (x-0≈x x)))
59+
y⁻¹ * (x⁻¹ * (x * y)) ≈⟨ *-congˡ (*-assoc x⁻¹ x y)
60+
y⁻¹ * ((x⁻¹ * x) * y) ≈⟨ *-congˡ (*-congʳ (*-congˡ (x-0≈x x)))
6161
y⁻¹ * ((x⁻¹ * (x - 0#)) * y) ≈⟨ *-congˡ (*-congʳ x⁻¹*x≈1) ⟩
6262
y⁻¹ * (1# * y) ≈⟨ *-congˡ (*-identityˡ y) ⟩
63-
y⁻¹ * y ≈˘⟨ *-congˡ (x-0≈x y)
63+
y⁻¹ * y ≈⟨ *-congˡ (x-0≈x y)
6464
y⁻¹ * (y - 0#) ≈⟨ y⁻¹*y≈1 ⟩
6565
1# ∎
6666

@@ -75,15 +75,15 @@ x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0)
7575

7676
y-x≈-[x-y] : y - x ≈ - (x - y)
7777
y-x≈-[x-y] = begin
78-
y - x ≈˘⟨ +-congʳ (-‿involutive y)
79-
- - y - x ≈˘⟨ -‿anti-homo-+ x (- y)
78+
y - x ≈⟨ +-congʳ (-‿involutive y)
79+
- - y - x ≈⟨ -‿anti-homo-+ x (- y)
8080
- (x - y) ∎
8181

8282
x-y⁻¹*y-x≈1 : (- x-y⁻¹) * (y - x) ≈ 1#
8383
x-y⁻¹*y-x≈1 = begin
84-
(- x-y⁻¹) * (y - x) ≈˘⟨ -‿distribˡ-* x-y⁻¹ (y - x)
84+
(- x-y⁻¹) * (y - x) ≈⟨ -‿distribˡ-* x-y⁻¹ (y - x)
8585
- (x-y⁻¹ * (y - x)) ≈⟨ -‿cong (*-congˡ y-x≈-[x-y]) ⟩
86-
- (x-y⁻¹ * - (x - y)) ≈˘⟨ -‿cong (-‿distribʳ-* x-y⁻¹ (x - y))
86+
- (x-y⁻¹ * - (x - y)) ≈⟨ -‿cong (-‿distribʳ-* x-y⁻¹ (x - y))
8787
- - (x-y⁻¹ * (x - y)) ≈⟨ -‿involutive (x-y⁻¹ * ((x - y))) ⟩
8888
x-y⁻¹ * (x - y) ≈⟨ InvX-Y .proj₂ .proj₁ ⟩
8989
1# ∎
@@ -100,7 +100,7 @@ x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0)
100100

101101
x-z⁻¹*y-z≈1 : x-z⁻¹ * (y - z) ≈ 1#
102102
x-z⁻¹*y-z≈1 = begin
103-
x-z⁻¹ * (y - z) ≈˘⟨ *-congˡ (+-congʳ x≈y)
103+
x-z⁻¹ * (y - z) ≈⟨ *-congˡ (+-congʳ x≈y)
104104
x-z⁻¹ * (x - z) ≈⟨ x-z⁻¹*x-z≈1# ⟩
105105
1# ∎
106106

src/Algebra/Consequences/Setoid.agda

+5-5
Original file line numberDiff line numberDiff line change
@@ -88,7 +88,7 @@ module _ {f : Op₁ A} (self : SelfInverse f) where
8888

8989
selfInverse⇒injective : Injective _≈_ _≈_ f
9090
selfInverse⇒injective {x} {y} x≈y = begin
91-
x ≈˘⟨ self x≈y
91+
x ≈⟨ self x≈y
9292
f (f y) ≈⟨ selfInverse⇒involutive y ⟩
9393
y ∎
9494

@@ -267,12 +267,12 @@ module _ {_•_ _◦_ : Op₂ A}
267267
_◦_ DistributesOver _•_
268268
_•_ DistributesOverˡ _◦_
269269
distrib+absorbs⇒distribˡ •-absorbs-◦ ◦-absorbs-• (◦-distribˡ-• , ◦-distribʳ-•) x y z = begin
270-
x • (y ◦ z) ≈˘⟨ •-cong (•-absorbs-◦ _ _) refl
270+
x • (y ◦ z) ≈⟨ •-cong (•-absorbs-◦ _ _) refl
271271
(x • (x ◦ y)) • (y ◦ z) ≈⟨ •-cong (•-cong refl (◦-comm _ _)) refl ⟩
272272
(x • (y ◦ x)) • (y ◦ z) ≈⟨ •-assoc _ _ _ ⟩
273-
x • ((y ◦ x) • (y ◦ z)) ≈˘⟨ •-cong refl (◦-distribˡ-• _ _ _)
274-
x • (y ◦ (x • z)) ≈˘⟨ •-cong (◦-absorbs-• _ _) refl
275-
(x ◦ (x • z)) • (y ◦ (x • z)) ≈˘⟨ ◦-distribʳ-• _ _ _
273+
x • ((y ◦ x) • (y ◦ z)) ≈⟨ •-cong refl (◦-distribˡ-• _ _ _)
274+
x • (y ◦ (x • z)) ≈⟨ •-cong (◦-absorbs-• _ _) refl
275+
(x ◦ (x • z)) • (y ◦ (x • z)) ≈⟨ ◦-distribʳ-• _ _ _
276276
(x • y) ◦ (x • z) ∎
277277

278278
------------------------------------------------------------------------

src/Algebra/Construct/LexProduct/Inner.agda

+18-18
Original file line numberDiff line numberDiff line change
@@ -82,8 +82,8 @@ module NaturalOrder where
8282

8383
≤∙ˡ-trans : Associative _≈₁_ _∙_ (a ∙ b) ≈₁ b (b ∙ c) ≈₁ c (a ∙ c) ≈₁ c
8484
≤∙ˡ-trans {a} {b} {c} ∙-assoc ab≈b bc≈c = begin
85-
a ∙ c ≈˘⟨ ∙-congˡ bc≈c
86-
a ∙ (b ∙ c) ≈˘⟨ ∙-assoc a b c
85+
a ∙ c ≈⟨ ∙-congˡ bc≈c
86+
a ∙ (b ∙ c) ≈⟨ ∙-assoc a b c
8787
(a ∙ b) ∙ c ≈⟨ ∙-congʳ ab≈b ⟩
8888
b ∙ c ≈⟨ bc≈c ⟩
8989
c ∎
@@ -179,51 +179,51 @@ assoc ∙-assoc ∙-comm ∙-sel ◦-assoc a b c x y z
179179
... | no ab≉a | yes ab≈b | yes bc≈b | yes bc≈c = begin
180180
lex (a ∙ b) c y z ≈⟨ cong₁ ab≈b ⟩
181181
lex b c y z ≈⟨ case₃ bc≈b bc≈c ⟩
182-
y ◦ z ≈˘⟨ case₂ ab≉a ab≈b
183-
lex a b x (y ◦ z) ≈˘⟨ cong₂ bc≈b
182+
y ◦ z ≈⟨ case₂ ab≉a ab≈b
183+
lex a b x (y ◦ z) ≈⟨ cong₂ bc≈b
184184
lex a (b ∙ c) x (y ◦ z) ∎
185185
... | no ab≉a | yes ab≈b | yes bc≈b | no bc≉c = begin
186186
lex (a ∙ b) c y z ≈⟨ cong₁ ab≈b ⟩
187187
lex b c y z ≈⟨ case₁ bc≈b bc≉c ⟩
188-
y ≈˘⟨ case₂ ab≉a ab≈b
189-
lex a b x y ≈˘⟨ cong₂ bc≈b
188+
y ≈⟨ case₂ ab≉a ab≈b
189+
lex a b x y ≈⟨ cong₂ bc≈b
190190
lex a (b ∙ c) x y ∎
191191
... | yes ab≈a | yes ab≈b | yes bc≈b | no bc≉c = begin
192192
lex (a ∙ b) c (x ◦ y) z ≈⟨ cong₁ ab≈b ⟩
193193
lex b c (x ◦ y) z ≈⟨ case₁ bc≈b bc≉c ⟩
194-
x ◦ y ≈˘⟨ case₃ ab≈a ab≈b
195-
lex a b x y ≈˘⟨ cong₂ bc≈b
194+
x ◦ y ≈⟨ case₃ ab≈a ab≈b
195+
lex a b x y ≈⟨ cong₂ bc≈b
196196
lex a (b ∙ c) x y ∎
197197
... | yes ab≈a | yes ab≈b | yes bc≈b | yes bc≈c = begin
198198
lex (a ∙ b) c (x ◦ y) z ≈⟨ cong₁ ab≈b ⟩
199199
lex b c (x ◦ y) z ≈⟨ case₃ bc≈b bc≈c ⟩
200200
(x ◦ y) ◦ z ≈⟨ ◦-assoc x y z ⟩
201-
x ◦ (y ◦ z) ≈˘⟨ case₃ ab≈a ab≈b
202-
lex a b x (y ◦ z) ≈˘⟨ cong₂ bc≈b
201+
x ◦ (y ◦ z) ≈⟨ case₃ ab≈a ab≈b
202+
lex a b x (y ◦ z) ≈⟨ cong₂ bc≈b
203203
lex a (b ∙ c) x (y ◦ z) ∎
204204
... | yes ab≈a | yes ab≈b | no bc≉b | yes bc≈c = begin
205205
lex (a ∙ b) c (x ◦ y) z ≈⟨ cong₁ ab≈b ⟩
206206
lex b c (x ◦ y) z ≈⟨ case₂ bc≉b bc≈c ⟩
207-
z ≈˘⟨ case₂ bc≉b bc≈c
208-
lex b c x z ≈˘⟨ cong₁₂ (M.trans (M.sym ab≈a) ab≈b) bc≈c
207+
z ≈⟨ case₂ bc≉b bc≈c
208+
lex b c x z ≈⟨ cong₁₂ (M.trans (M.sym ab≈a) ab≈b) bc≈c
209209
lex a (b ∙ c) x z ∎
210210
... | yes ab≈a | no ab≉b | yes bc≈b | yes bc≈c = begin
211211
lex (a ∙ b) c x z ≈⟨ cong₁₂ ab≈a (M.trans (M.sym bc≈c) bc≈b) ⟩
212212
lex a b x z ≈⟨ case₁ ab≈a ab≉b ⟩
213-
x ≈˘⟨ case₁ ab≈a ab≉b
214-
lex a b x (y ◦ z) ≈˘⟨ cong₂ bc≈b
213+
x ≈⟨ case₁ ab≈a ab≉b
214+
lex a b x (y ◦ z) ≈⟨ cong₂ bc≈b
215215
lex a (b ∙ c) x (y ◦ z) ∎
216216
... | no ab≉a | yes ab≈b | no bc≉b | yes bc≈c = begin
217217
lex (a ∙ b) c y z ≈⟨ cong₁ ab≈b ⟩
218218
lex b c y z ≈⟨ case₂ bc≉b bc≈c ⟩
219-
z ≈˘⟨ uncurry′ case₂ (<∙ˡ-trans ∙-assoc ∙-comm ab≈b ab≉a bc≈c)
220-
lex a c x z ≈˘⟨ cong₂ bc≈c
219+
z ≈⟨ uncurry′ case₂ (<∙ˡ-trans ∙-assoc ∙-comm ab≈b ab≉a bc≈c)
220+
lex a c x z ≈⟨ cong₂ bc≈c
221221
lex a (b ∙ c) x z ∎
222222
... | yes ab≈a | no ab≉b | yes bc≈b | no bc≉c = begin
223223
lex (a ∙ b) c x z ≈⟨ cong₁ ab≈a ⟩
224224
lex a c x z ≈⟨ uncurry′ case₁ (<∙ʳ-trans ∙-assoc ∙-comm ab≈a bc≈b bc≉c) ⟩
225-
x ≈˘⟨ case₁ ab≈a ab≉b
226-
lex a b x y ≈˘⟨ cong₂ bc≈b
225+
x ≈⟨ case₁ ab≈a ab≉b
226+
lex a b x y ≈⟨ cong₂ bc≈b
227227
lex a (b ∙ c) x y ∎
228228

229229
comm : Commutative _≈₁_ _∙_ Commutative _≈₂_ _◦_

0 commit comments

Comments
 (0)