|
| 1 | +------------------------------------------------------------------------ |
| 2 | +-- The Agda standard library |
| 3 | +-- |
| 4 | +-- An equational reasoning library for propositional equality over |
| 5 | +-- vectors of different indices using cast. |
| 6 | +------------------------------------------------------------------------ |
| 7 | + |
| 8 | +{-# OPTIONS --cubical-compatible --safe #-} |
| 9 | + |
| 10 | +module README.Data.Vec.Relation.Binary.Equality.Cast where |
| 11 | + |
| 12 | +open import Agda.Primitive |
| 13 | +open import Data.List.Base as L using (List) |
| 14 | +import Data.List.Properties as Lₚ |
| 15 | +open import Data.Nat.Base |
| 16 | +open import Data.Nat.Properties |
| 17 | +open import Data.Vec.Base |
| 18 | +open import Data.Vec.Properties |
| 19 | +open import Data.Vec.Relation.Binary.Equality.Cast |
| 20 | +open import Relation.Binary.PropositionalEquality |
| 21 | + using (_≡_; refl; trans; sym; cong; subst; module ≡-Reasoning) |
| 22 | + |
| 23 | +private variable |
| 24 | + a : Level |
| 25 | + A : Set a |
| 26 | + l m n o : ℕ |
| 27 | + xs ys zs ws : Vec A n |
| 28 | + |
| 29 | + |
| 30 | +-- To see example usages of this library, scroll to the combinators |
| 31 | +-- section. |
| 32 | + |
| 33 | + |
| 34 | +------------------------------------------------------------------------ |
| 35 | +-- Motivation |
| 36 | +-- |
| 37 | +-- The `cast` function is the computational variant of `subst` for |
| 38 | +-- vectors. Since `cast` computes under vector constructors, it |
| 39 | +-- enables reasoning about vectors with non-definitionally equal indices |
| 40 | +-- by induction. See, e.g., Jacques Carette's comment in issue #1668. |
| 41 | +-- <https://github.com/agda/agda-stdlib/pull/1668#issuecomment-1003449509> |
| 42 | +-- |
| 43 | +-- Suppose we want to prove that ‘xs ++ [] ≡ xs’. Because `xs ++ []` |
| 44 | +-- has type `Vec A (n + 0)` while `xs` has type `Vec A n`, they cannot |
| 45 | +-- be directly related by homogeneous equality. |
| 46 | +-- To resolve the issue, `++-right-identity` uses `cast` to recast |
| 47 | +-- `xs ++ []` as a vector in `Vec A n`. |
| 48 | +-- |
| 49 | +++-right-identity : ∀ .(eq : n + 0 ≡ n) (xs : Vec A n) → cast eq (xs ++ []) ≡ xs |
| 50 | +++-right-identity eq [] = refl |
| 51 | +++-right-identity eq (x ∷ xs) = cong (x ∷_) (++-right-identity (cong pred eq) xs) |
| 52 | +-- |
| 53 | +-- When the input is `x ∷ xs`, because `cast eq (x ∷ _)` equals |
| 54 | +-- `x ∷ cast (cong pred eq) _`, the proof obligation |
| 55 | +-- cast eq (x ∷ xs ++ []) ≡ x ∷ xs |
| 56 | +-- simplifies to |
| 57 | +-- x :: cast (cong pred eq) (xs ++ []) ≡ x ∷ xs |
| 58 | + |
| 59 | + |
| 60 | +-- Although `cast` makes it possible to prove vector identities by ind- |
| 61 | +-- uction, the explicit type-casting nature poses a significant barrier |
| 62 | +-- to code reuse in larger proofs. For example, consider the identity |
| 63 | +-- ‘fromList (xs L.∷ʳ x) ≡ (fromList xs) ∷ʳ x’ where `L._∷ʳ_` is the |
| 64 | +-- snoc function of lists. We have |
| 65 | +-- |
| 66 | +-- fromList (xs L.∷ʳ x) : Vec A (L.length (xs L.∷ʳ x)) |
| 67 | +-- = {- by definition -} |
| 68 | +-- fromList (xs L.++ L.[ x ]) : Vec A (L.length (xs L.++ L.[ x ])) |
| 69 | +-- = {- by fromList-++ -} |
| 70 | +-- fromList xs ++ fromList L.[ x ] : Vec A (L.length xs + L.length [ x ]) |
| 71 | +-- = {- by definition -} |
| 72 | +-- fromList xs ++ [ x ] : Vec A (L.length xs + 1) |
| 73 | +-- = {- by unfold-∷ʳ -} |
| 74 | +-- fromList xs ∷ʳ x : Vec A (suc (L.length xs)) |
| 75 | +-- where |
| 76 | +-- fromList-++ : cast _ (fromList (xs L.++ ys)) ≡ fromList xs ++ fromList ys |
| 77 | +-- unfold-∷ʳ : cast _ (xs ∷ʳ x) ≡ xs ++ [ x ] |
| 78 | +-- |
| 79 | +-- Although the identity itself is simple, the reasoning process changes |
| 80 | +-- the index in the type twice. Consequently, its Agda translation must |
| 81 | +-- insert two `cast`s in the proof. Moreover, the proof first has to |
| 82 | +-- rearrange (the Agda version of) the identity into one with two |
| 83 | +-- `cast`s, resulting in lots of boilerplate code as demonstrated by |
| 84 | +-- `example1a-fromList-∷ʳ`. |
| 85 | +example1a-fromList-∷ʳ : ∀ (x : A) xs .(eq : L.length (xs L.∷ʳ x) ≡ suc (L.length xs)) → |
| 86 | + cast eq (fromList (xs L.∷ʳ x)) ≡ fromList xs ∷ʳ x |
| 87 | +example1a-fromList-∷ʳ x xs eq = begin |
| 88 | + cast eq (fromList (xs L.∷ʳ x)) ≡⟨⟩ |
| 89 | + cast eq (fromList (xs L.++ L.[ x ])) ≡˘⟨ cast-trans eq₁ eq₂ (fromList (xs L.++ L.[ x ])) ⟩ |
| 90 | + cast eq₂ (cast eq₁ (fromList (xs L.++ L.[ x ]))) ≡⟨ cong (cast eq₂) (fromList-++ xs) ⟩ |
| 91 | + cast eq₂ (fromList xs ++ [ x ]) ≡⟨ ≈-sym (unfold-∷ʳ (sym eq₂) x (fromList xs)) ⟩ |
| 92 | + fromList xs ∷ʳ x ∎ |
| 93 | + where |
| 94 | + open ≡-Reasoning |
| 95 | + eq₁ = Lₚ.length-++ xs {L.[ x ]} |
| 96 | + eq₂ = +-comm (L.length xs) 1 |
| 97 | + |
| 98 | +-- The `cast`s are irrelevant to core of the proof. At the same time, |
| 99 | +-- they can be inferred from the lemmas used during the reasoning steps |
| 100 | +-- (e.g. `fromList-++` and `unfold-∷ʳ`). To eliminate the boilerplate, |
| 101 | +-- this library provides a set of equational reasoning combinators for |
| 102 | +-- equality of the form `cast eq xs ≡ ys`. |
| 103 | +example1b-fromList-∷ʳ : ∀ (x : A) xs .(eq : L.length (xs L.∷ʳ x) ≡ suc (L.length xs)) → |
| 104 | + cast eq (fromList (xs L.∷ʳ x)) ≡ fromList xs ∷ʳ x |
| 105 | +example1b-fromList-∷ʳ x xs eq = begin |
| 106 | + fromList (xs L.∷ʳ x) ≈⟨⟩ |
| 107 | + fromList (xs L.++ L.[ x ]) ≈⟨ fromList-++ xs ⟩ |
| 108 | + fromList xs ++ [ x ] ≈˘⟨ unfold-∷ʳ (+-comm 1 (L.length xs)) x (fromList xs) ⟩ |
| 109 | + fromList xs ∷ʳ x ∎ |
| 110 | + where open CastReasoning |
| 111 | + |
| 112 | + |
| 113 | +------------------------------------------------------------------------ |
| 114 | +-- Combinators |
| 115 | +-- |
| 116 | +-- Let `xs ≈[ m≡n ] ys` denote `cast m≡n xs ≡ ys`. We have reflexivity, |
| 117 | +-- symmetry and transitivity: |
| 118 | +-- ≈-reflexive : xs ≈[ refl ] xs |
| 119 | +-- ≈-sym : xs ≈[ m≡n ] ys → ys ≈[ sym m≡n ] xs |
| 120 | +-- ≈-trans : xs ≈[ m≡n ] ys → ys ≈[ n≡o ] zs → xs ≈[ trans m≡n n≡o ] zs |
| 121 | +-- Accordingly, `_≈[_]_` admits the standard set of equational reasoning |
| 122 | +-- combinators. Suppose `≈-eqn : xs ≈[ m≡n ] ys`, |
| 123 | +-- xs ≈⟨ ≈-eqn ⟩ -- `_≈⟨_⟩_` takes a `_≈[_]_` step, adjusting |
| 124 | +-- ys -- the index at the same time |
| 125 | +-- |
| 126 | +-- ys ≈˘⟨ ≈-eqn ⟩ -- `_≈˘⟨_⟩_` takes a symmetric `_≈[_]_` step |
| 127 | +-- xs |
| 128 | +example2a : ∀ .(eq : suc m + n ≡ m + suc n) (xs : Vec A m) a ys → |
| 129 | + cast eq ((reverse xs ∷ʳ a) ++ ys) ≡ reverse xs ++ (a ∷ ys) |
| 130 | +example2a eq xs a ys = begin |
| 131 | + (reverse xs ∷ʳ a) ++ ys ≈⟨ ∷ʳ-++ eq a (reverse xs) ⟩ -- index: suc m + n |
| 132 | + reverse xs ++ (a ∷ ys) ∎ -- index: m + suc n |
| 133 | + where open CastReasoning |
| 134 | + |
| 135 | +-- To interoperate with `_≡_`, this library provides `_≂⟨_⟩_` (\-~) for |
| 136 | +-- taking a `_≡_` step during equational reasoning. |
| 137 | +-- Let `≡-eqn : xs ≡ ys`, then |
| 138 | +-- xs ≂⟨ ≡-eqn ⟩ -- Takes a `_≡_` step; no change to the index |
| 139 | +-- ys |
| 140 | +-- |
| 141 | +-- ys ≂˘⟨ ≡-eqn ⟩ -- Takes a symmetric `_≡_` step |
| 142 | +-- xs |
| 143 | +-- Equivalently, `≈-reflexive` injects `_≡_` into `_≈[_]_`. That is, |
| 144 | +-- `xs ≂⟨ ≡-eqn ⟩ ys` is the same as `xs ≈⟨ ≈-reflexive ≡-eqn ⟩ ys`. |
| 145 | +-- Extending `example2a`, we have: |
| 146 | +example2b : ∀ .(eq : suc m + n ≡ m + suc n) (xs : Vec A m) a ys → |
| 147 | + cast eq ((a ∷ xs) ʳ++ ys) ≡ xs ʳ++ (a ∷ ys) |
| 148 | +example2b eq xs a ys = begin |
| 149 | + (a ∷ xs) ʳ++ ys ≂⟨ unfold-ʳ++ (a ∷ xs) ys ⟩ -- index: suc m + n |
| 150 | + reverse (a ∷ xs) ++ ys ≂⟨ cong (_++ ys) (reverse-∷ a xs) ⟩ -- index: suc m + n |
| 151 | + (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 |
| 153 | + xs ʳ++ (a ∷ ys) ∎ -- index: m + suc n |
| 154 | + where open CastReasoning |
| 155 | + |
| 156 | +-- Oftentimes index-changing identities apply to only part of the proof |
| 157 | +-- term. When reasoning about `_≡_`, `cong` shifts the focus to the |
| 158 | +-- subterm of interest. In this library, `≈-cong` does a similar job. |
| 159 | +-- Suppose `f : A → B`, `xs : B`, `ys zs : A`, `ys≈zs : ys ≈[ _ ] zs` |
| 160 | +-- and `xs≈f⟨c·ys⟩ : xs ≈[ _ ] f (cast _ ys)`, we have |
| 161 | +-- xs ≈⟨ ≈-cong f xs≈f⟨c·ys⟩ ys≈zs ⟩ |
| 162 | +-- f zs |
| 163 | +-- The reason for having the extra argument `xs≈f⟨c·ys⟩` is to expose |
| 164 | +-- `cast` in the subterm in order to apply the step `ys≈zs`. When using |
| 165 | +-- ordinary `cong` the proof has to explicitly push `cast` inside: |
| 166 | +-- xs ≈⟨ xs≈f⟨c·ys⟩ ⟩ |
| 167 | +-- f (cast _ ys) ≂⟨ cong f ys≈zs ⟩ |
| 168 | +-- f zs |
| 169 | +-- Note. Technically, `A` and `B` should be vectors of different length |
| 170 | +-- and that `ys`, `zs` are vectors of non-definitionally equal index. |
| 171 | +example3a-fromList-++-++ : {xs ys zs : List A} → |
| 172 | + .(eq : L.length (xs L.++ ys L.++ zs) ≡ |
| 173 | + L.length xs + (L.length ys + L.length zs)) → |
| 174 | + cast eq (fromList (xs L.++ ys L.++ zs)) ≡ |
| 175 | + fromList xs ++ fromList ys ++ fromList zs |
| 176 | +example3a-fromList-++-++ {xs = xs} {ys} {zs} eq = begin |
| 177 | + fromList (xs L.++ ys L.++ zs) ≈⟨ fromList-++ xs ⟩ |
| 178 | + fromList xs ++ fromList (ys L.++ zs) ≈⟨ ≈-cong (fromList xs ++_) (cast-++ʳ (Lₚ.length-++ ys) (fromList xs)) |
| 179 | + (fromList-++ ys) ⟩ |
| 180 | + fromList xs ++ fromList ys ++ fromList zs ∎ |
| 181 | + where open CastReasoning |
| 182 | + |
| 183 | +-- As an alternative, one can manually apply `cast-++ʳ` to expose `cast` |
| 184 | +-- in the subterm. However, this unavoidably duplicates the proof term. |
| 185 | +example3b-fromList-++-++′ : {xs ys zs : List A} → |
| 186 | + .(eq : L.length (xs L.++ ys L.++ zs) ≡ |
| 187 | + L.length xs + (L.length ys + L.length zs)) → |
| 188 | + cast eq (fromList (xs L.++ ys L.++ zs)) ≡ |
| 189 | + fromList xs ++ fromList ys ++ fromList zs |
| 190 | +example3b-fromList-++-++′ {xs = xs} {ys} {zs} eq = begin |
| 191 | + fromList (xs L.++ ys L.++ zs) ≈⟨ fromList-++ xs ⟩ |
| 192 | + fromList xs ++ fromList (ys L.++ zs) ≈⟨ cast-++ʳ (Lₚ.length-++ ys) (fromList xs) ⟩ |
| 193 | + fromList xs ++ cast _ (fromList (ys L.++ zs)) ≂⟨ cong (fromList xs ++_) (fromList-++ ys) ⟩ |
| 194 | + fromList xs ++ fromList ys ++ fromList zs ∎ |
| 195 | + where open CastReasoning |
| 196 | + |
| 197 | +-- `≈-cong` can be chained together much like how `cong` can be nested. |
| 198 | +-- In this example, `unfold-∷ʳ` is applied to the term `xs ++ [ a ]` |
| 199 | +-- in `(_++ ys)` inside of `reverse`. Thus the proof employs two |
| 200 | +-- `≈-cong`. |
| 201 | +example4-cong² : ∀ .(eq : (m + 1) + n ≡ n + suc m) a (xs : Vec A m) ys → |
| 202 | + cast eq (reverse ((xs ++ [ a ]) ++ ys)) ≡ ys ʳ++ reverse (xs ∷ʳ a) |
| 203 | +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)) |
| 205 | + (≈-cong (_++ ys) (cast-++ˡ (+-comm 1 m) (xs ∷ʳ a)) |
| 206 | + (unfold-∷ʳ _ a xs)) ⟩ |
| 207 | + reverse ((xs ∷ʳ a) ++ ys) ≈⟨ reverse-++ (+-comm (suc m) n) (xs ∷ʳ a) ys ⟩ |
| 208 | + reverse ys ++ reverse (xs ∷ʳ a) ≂˘⟨ unfold-ʳ++ ys (reverse (xs ∷ʳ a)) ⟩ |
| 209 | + ys ʳ++ reverse (xs ∷ʳ a) ∎ |
| 210 | + where open CastReasoning |
| 211 | + |
| 212 | +------------------------------------------------------------------------ |
| 213 | +-- Interoperation between `_≈[_]_` and `_≡_` |
| 214 | +-- |
| 215 | +-- This library is designed to interoperate with `_≡_`. Examples in the |
| 216 | +-- combinators section showed how to apply `_≂⟨_⟩_` to take an `_≡_` |
| 217 | +-- step during equational reasoning about `_≈[_]_`. Recall that |
| 218 | +-- `xs ≈[ m≡n ] ys` is a shorthand for `cast m≡n xs ≡ ys`, the |
| 219 | +-- combinator is essentially the composition of `_≡_` on the left-hand |
| 220 | +-- side of `_≈[_]_`. Dually, the combinator `_≃⟨_⟩_` composes `_≡_` on |
| 221 | +-- the right-hand side of `_≈[_]_`. Thus `_≃⟨_⟩_` intuitively ends the |
| 222 | +-- reasoning system of `_≈[_]_` and switches back to the reasoning |
| 223 | +-- system of `_≡_`. |
| 224 | +example5-fromList-++-++′ : {xs ys zs : List A} → |
| 225 | + .(eq : L.length (xs L.++ ys L.++ zs) ≡ |
| 226 | + L.length xs + (L.length ys + L.length zs)) → |
| 227 | + cast eq (fromList (xs L.++ ys L.++ zs)) ≡ |
| 228 | + fromList xs ++ fromList ys ++ fromList zs |
| 229 | +example5-fromList-++-++′ {xs = xs} {ys} {zs} eq = begin |
| 230 | + fromList (xs L.++ ys L.++ zs) ≈⟨ fromList-++ xs ⟩ |
| 231 | + fromList xs ++ fromList (ys L.++ zs) ≃⟨ cast-++ʳ (Lₚ.length-++ ys) (fromList xs) ⟩ |
| 232 | + fromList xs ++ cast _ (fromList (ys L.++ zs)) ≡⟨ cong (fromList xs ++_) (fromList-++ ys) ⟩ |
| 233 | + fromList xs ++ fromList ys ++ fromList zs ≡-∎ |
| 234 | + where open CastReasoning |
| 235 | + |
| 236 | +-- Of course, it is possible to start with the reasoning system of `_≡_` |
| 237 | +-- and then switch to the reasoning system of `_≈[_]_`. |
| 238 | +example6a-reverse-∷ʳ : ∀ x (xs : Vec A n) → reverse (xs ∷ʳ x) ≡ x ∷ reverse xs |
| 239 | +example6a-reverse-∷ʳ {n = n} x xs = begin-≡ |
| 240 | + reverse (xs ∷ʳ x) ≡˘⟨ ≈-reflexive refl ⟩ |
| 241 | + reverse (xs ∷ʳ x) ≈⟨ ≈-cong reverse (cast-reverse _ _) (unfold-∷ʳ (+-comm 1 n) x xs) ⟩ |
| 242 | + reverse (xs ++ [ x ]) ≈⟨ reverse-++ (+-comm n 1) xs [ x ] ⟩ |
| 243 | + x ∷ reverse xs ∎ |
| 244 | + where open CastReasoning |
| 245 | + |
| 246 | +example6b-reverse-∷ʳ-by-induction : ∀ x (xs : Vec A n) → reverse (xs ∷ʳ x) ≡ x ∷ reverse xs |
| 247 | +example6b-reverse-∷ʳ-by-induction x [] = refl |
| 248 | +example6b-reverse-∷ʳ-by-induction x (y ∷ xs) = begin |
| 249 | + reverse (y ∷ (xs ∷ʳ x)) ≡⟨ reverse-∷ y (xs ∷ʳ x) ⟩ |
| 250 | + reverse (xs ∷ʳ x) ∷ʳ y ≡⟨ cong (_∷ʳ y) (example6b-reverse-∷ʳ-by-induction x xs) ⟩ |
| 251 | + (x ∷ reverse xs) ∷ʳ y ≡⟨⟩ |
| 252 | + x ∷ (reverse xs ∷ʳ y) ≡˘⟨ cong (x ∷_) (reverse-∷ y xs) ⟩ |
| 253 | + x ∷ reverse (y ∷ xs) ∎ |
| 254 | + where open ≡-Reasoning |
0 commit comments