Skip to content

Commit cc6addb

Browse files
committed
move ≈-cong′ to Cast
1 parent 438f95a commit cc6addb

File tree

2 files changed

+10
-9
lines changed

2 files changed

+10
-9
lines changed

src/Data/Vec/Properties.agda

Lines changed: 1 addition & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -369,7 +369,7 @@ lookup∘update′ {i = i} {j} i≢j xs y = lookup∘updateAt′ i j i≢j xs
369369
-- cast
370370

371371
open VecCast public
372-
using (cast-is-id; cast-trans)
372+
using (cast-is-id; cast-trans; ≈-cong′)
373373

374374
subst-is-cast : (eq : m ≡ n) (xs : Vec A m) subst (Vec A) eq xs ≡ cast eq xs
375375
subst-is-cast refl xs = sym (cast-is-id refl xs)
@@ -381,12 +381,6 @@ cast-sym eq {xs = x ∷ xs} {ys = y ∷ ys} xxs[eq]≡yys =
381381
let x≡y , xs[eq]≡ys = ∷-injective xxs[eq]≡yys
382382
in cong₂ _∷_ (sym x≡y) (cast-sym (suc-injective eq) xs[eq]≡ys)
383383

384-
≈-cong′ : {f-len : ℕ} (f : {n} Vec A n Vec B (f-len n))
385-
{m n} {xs : Vec A m} {ys : Vec A n} .{eq} xs ≈[ eq ] ys
386-
f xs ≈[ cong f-len eq ] f ys
387-
≈-cong′ f {m = zero} {n = zero} {xs = []} {ys = []} refl = cast-is-id refl (f [])
388-
≈-cong′ f {m = suc m} {n = suc n} {xs = x ∷ xs} {ys = y ∷ ys} refl = ≈-cong′ (f ∘ (x ∷_)) refl
389-
390384
------------------------------------------------------------------------
391385
-- map
392386

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

Lines changed: 9 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -13,6 +13,7 @@
1313
module Data.Vec.Relation.Binary.Equality.Cast where
1414

1515
open import Level using (Level)
16+
open import Function using (_∘_)
1617
open import Data.Nat.Base using (ℕ; zero; suc)
1718
open import Data.Nat.Properties using (suc-injective)
1819
open import Data.Vec.Base
@@ -25,8 +26,8 @@ open import Relation.Binary.PropositionalEquality.Properties
2526

2627
private
2728
variable
28-
a : Level
29-
A : Set a
29+
a b : Level
30+
A B : Set a
3031
l m n o :
3132
xs ys zs : Vec A n
3233

@@ -70,6 +71,12 @@ xs ≈[ eq ] ys = cast eq xs ≡ ys
7071
zs ∎
7172
where open ≡-Reasoning
7273

74+
≈-cong′ : {f-len : ℕ} (f : {n} Vec A n Vec B (f-len n))
75+
{m n} {xs : Vec A m} {ys : Vec A n} .{eq} xs ≈[ eq ] ys
76+
f xs ≈[ cong f-len eq ] f ys
77+
≈-cong′ f {m = zero} {n = zero} {xs = []} {ys = []} refl = cast-is-id refl (f [])
78+
≈-cong′ f {m = suc m} {n = suc n} {xs = x ∷ xs} {ys = y ∷ ys} refl = ≈-cong′ (f ∘ (x ∷_)) refl
79+
7380
------------------------------------------------------------------------
7481
-- Reasoning combinators
7582

0 commit comments

Comments
 (0)