Skip to content

Commit 4fdaf7c

Browse files
committed
fixes agda#1131
1 parent 606bea8 commit 4fdaf7c

File tree

2 files changed

+17
-3
lines changed

2 files changed

+17
-3
lines changed

src/Data/List/Relation/Binary/Equality/Setoid.agda

+6
Original file line numberDiff line numberDiff line change
@@ -116,6 +116,12 @@ foldr⁺ ∙⇔◦ e≈f xs≋ys = PW.foldr⁺ ∙⇔◦ e≈f xs≋ys
116116
++⁺ : {ws xs ys zs} ws ≋ xs ys ≋ zs ws ++ ys ≋ xs ++ zs
117117
++⁺ = PW.++⁺
118118

119+
++⁺ʳ : xs {ys zs} ys ≋ zs xs ++ ys ≋ xs ++ zs
120+
++⁺ʳ xs = PW.++⁺ʳ refl xs
121+
122+
++⁺ˡ : {ws xs} ws ≋ xs zs ws ++ zs ≋ xs ++ zs
123+
++⁺ˡ rs = PW.++⁺ˡ refl rs
124+
119125
++-cancelˡ : xs {ys zs} xs ++ ys ≋ xs ++ zs ys ≋ zs
120126
++-cancelˡ xs = PW.++-cancelˡ xs
121127

src/Data/List/Relation/Binary/Pointwise.agda

+11-3
Original file line numberDiff line numberDiff line change
@@ -26,7 +26,7 @@ open import Relation.Nullary hiding (Irrelevant)
2626
import Relation.Nullary.Decidable as Dec using (map′)
2727
open import Relation.Unary as U using (Pred)
2828
open import Relation.Binary.Core renaming (Rel to Rel₂)
29-
open import Relation.Binary.Definitions using (_Respects_; _Respects₂_)
29+
open import Relation.Binary.Definitions using (Reflexive;_Respects_; _Respects₂_)
3030
open import Relation.Binary.Bundles using (Setoid; DecSetoid; Preorder; Poset)
3131
open import Relation.Binary.Structures using (IsEquivalence; IsDecEquivalence; IsPartialOrder; IsPreorder)
3232
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
@@ -166,6 +166,15 @@ tabulate⁻ {n = suc n} (x∼y ∷ xs∼ys) (fsuc i) = tabulate⁻ xs∼ys i
166166
++-cancelʳ {xs = xs} (y ∷ ys) [] eq =
167167
contradiction (≡.trans (≡.sym (length-++ (y ∷ ys))) (Pointwise-length eq)) (m≢1+n+m (length xs) ∘ ≡.sym)
168168

169+
module _ (rfl : Reflexive R) where
170+
171+
++⁺ʳ : xs Pointwise R ys zs Pointwise R (xs ++ ys) (xs ++ zs)
172+
++⁺ʳ xs = ++⁺ (refl rfl)
173+
174+
++⁺ˡ : Pointwise R ws xs zs Pointwise R (ws ++ zs) (xs ++ zs)
175+
++⁺ˡ rs zs = ++⁺ rs (refl rfl)
176+
177+
169178
------------------------------------------------------------------------
170179
-- concat
171180

@@ -261,8 +270,7 @@ lookup⁺ (_ ∷ Rxys) (fsuc i) = lookup⁺ Rxys i
261270

262271
Pointwise-≡⇒≡ : Pointwise {A = A} _≡_ ⇒ _≡_
263272
Pointwise-≡⇒≡ [] = ≡.refl
264-
Pointwise-≡⇒≡ (≡.refl ∷ xs∼ys) with Pointwise-≡⇒≡ xs∼ys
265-
... | ≡.refl = ≡.refl
273+
Pointwise-≡⇒≡ (≡.refl ∷ xs∼ys) = ≡.cong (_ ∷_) (Pointwise-≡⇒≡ xs∼ys)
266274

267275
≡⇒Pointwise-≡ : _≡_ ⇒ Pointwise {A = A} _≡_
268276
≡⇒Pointwise-≡ ≡.refl = refl ≡.refl

0 commit comments

Comments
 (0)