Skip to content

Commit 0734397

Browse files
Saransh-cppMatthewDaggitt
authored andcommitted
Sync iterate and replicate for List and Vec (#2051)
1 parent 5473bb8 commit 0734397

File tree

6 files changed

+102
-18
lines changed

6 files changed

+102
-18
lines changed

CHANGELOG.md

+27-3
Original file line numberDiff line numberDiff line change
@@ -1058,27 +1058,37 @@ Non-backwards compatible changes
10581058
Tactic.RingSolver
10591059
Tactic.RingSolver.Core.NatSet
10601060
```
1061+
10611062
* Moved & renamed from `Data.Vec.Relation.Unary.All`
10621063
to `Data.Vec.Relation.Unary.All.Properties`:
10631064
```
10641065
lookup ↦ lookup⁺
10651066
tabulate ↦ lookup⁻
10661067
```
1068+
10671069
* Renamed in `Data.Vec.Relation.Unary.Linked.Properties`
10681070
and `Codata.Guarded.Stream.Relation.Binary.Pointwise`:
10691071
```
10701072
lookup ↦ lookup⁺
10711073
```
1074+
10721075
* Added the following new definitions to `Data.Vec.Relation.Unary.All`:
10731076
```
10741077
lookupAny : All P xs → (i : Any Q xs) → (P ∩ Q) (Any.lookup i)
10751078
lookupWith : ∀[ P ⇒ Q ⇒ R ] → All P xs → (i : Any Q xs) → R (Any.lookup i)
10761079
lookup : All P xs → (∀ {x} → x ∈ₚ xs → P x)
10771080
lookupₛ : P Respects _≈_ → All P xs → (∀ {x} → x ∈ xs → P x)
10781081
```
1082+
10791083
* `excluded-middle` in `Relation.Nullary.Decidable.Core` has been renamed to
10801084
`¬¬-excluded-middle`.
10811085
1086+
* `iterate` in `Data.Vec.Base` now takes `n` (the length of `Vec`) as an
1087+
explicit argument.
1088+
```agda
1089+
iterate : (A → A) → A → ∀ n → Vec A n
1090+
```
1091+
10821092
Major improvements
10831093
------------------
10841094
@@ -2441,7 +2451,7 @@ Additions to existing modules
24412451
gcd-zero : Zero 1ℤ gcd
24422452
```
24432453

2444-
* Added new functions in `Data.List.Base`:
2454+
* Added new functions and definitions to `Data.List.Base`:
24452455
```agda
24462456
takeWhileᵇ : (A → Bool) → List A → List A
24472457
dropWhileᵇ : (A → Bool) → List A → List A
@@ -2466,6 +2476,7 @@ Additions to existing modules
24662476
++-rawMagma : Set a → RawMagma a _
24672477
++-[]-rawMonoid : Set a → RawMonoid a _
24682478
2479+
iterate : (A → A) → A → ℕ → List A
24692480
insertAt : (xs : List A) → Fin (suc (length xs)) → A → List A
24702481
updateAt : (xs : List A) → Fin (length xs) → (A → A) → List A
24712482
```
@@ -2527,14 +2538,21 @@ Additions to existing modules
25272538
drop-take-suc-tabulate : drop m (take (suc m) (tabulate f)) ≡ [ f i ]
25282539
25292540
take-all : n ≥ length xs → take n xs ≡ xs
2541+
drop-all : n ≥ length xs → drop n xs ≡ []
25302542
25312543
take-[] : take m [] ≡ []
25322544
drop-[] : drop m [] ≡ []
25332545
2534-
map-replicate : map f (replicate n x) ≡ replicate n (f x)
2535-
25362546
drop-drop : drop n (drop m xs) ≡ drop (m + n) xs
25372547
2548+
lookup-replicate : lookup (replicate n x) i ≡ x
2549+
map-replicate : map f (replicate n x) ≡ replicate n (f x)
2550+
zipWith-replicate : zipWith _⊕_ (replicate n x) (replicate n y) ≡ replicate n (x ⊕ y)
2551+
2552+
length-iterate : length (iterate f x n) ≡ n
2553+
iterate-id : iterate id x n ≡ replicate n x
2554+
lookup-iterate : lookup (iterate f x n) (cast (sym (length-iterate f x n)) i) ≡ ℕ.iterate f x (toℕ i)
2555+
25382556
length-insertAt : length (insertAt xs i v) ≡ suc (length xs)
25392557
length-removeAt′ : length xs ≡ suc (length (removeAt xs k))
25402558
removeAt-insertAt : removeAt (insertAt xs i v) ((cast (sym (length-insertAt xs i v)) i)) ≡ xs
@@ -2964,6 +2982,12 @@ Additions to existing modules
29642982
lookup-cast₂ : lookup xs (Fin.cast eq i) ≡ lookup (cast (sym eq) xs) i
29652983
cast-reverse : cast eq ∘ reverse ≗ reverse ∘ cast eq
29662984
2985+
iterate-id : iterate id x n ≡ replicate x
2986+
take-iterate : take n (iterate f x (n + m)) ≡ iterate f x n
2987+
drop-iterate : drop n (iterate f x n) ≡ []
2988+
lookup-iterate : lookup (iterate f x n) i ≡ ℕ.iterate f x (toℕ i)
2989+
toList-iterate : toList (iterate f x n) ≡ List.iterate f x n
2990+
29672991
zipwith-++ : zipWith f (xs ++ ys) (xs' ++ ys') ≡ zipWith f xs xs' ++ zipWith f ys ys'
29682992
29692993
++-assoc : cast eq ((xs ++ ys) ++ zs) ≡ xs ++ (ys ++ zs)

src/Codata/Guarded/Stream/Properties.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -255,7 +255,7 @@ lookup-interleave-odd (suc n) as bs = lookup-interleave-odd n (as .tail) (bs .ta
255255
------------------------------------------------------------------------
256256
-- Properties of take
257257

258-
take-iterate : n f (x : A) take n (iterate f x) ≡ Vec.iterate f x
258+
take-iterate : n f (x : A) take n (iterate f x) ≡ Vec.iterate f x n
259259
take-iterate zero f x = P.refl
260260
take-iterate (suc n) f x = cong (x ∷_) (take-iterate n f (f x))
261261

src/Data/List/Base.agda

+4
Original file line numberDiff line numberDiff line change
@@ -188,6 +188,10 @@ replicate : ℕ → A → List A
188188
replicate zero x = []
189189
replicate (suc n) x = x ∷ replicate n x
190190

191+
iterate : (A A) A List A
192+
iterate f e zero = []
193+
iterate f e (suc n) = e ∷ iterate f (f e) n
194+
191195
inits : List A List (List A)
192196
inits [] = [] ∷ []
193197
inits (x ∷ xs) = [] ∷ map (x ∷_) (inits xs)

src/Data/List/Properties.agda

+43-11
Original file line numberDiff line numberDiff line change
@@ -43,6 +43,7 @@ open import Relation.Nullary using (¬_; Dec; does; _because_; yes; no; contradi
4343
open import Relation.Nullary.Decidable as Decidable using (isYes; map′; ⌊_⌋; ¬?; _×-dec_)
4444
open import Relation.Unary using (Pred; Decidable; ∁)
4545
open import Relation.Unary.Properties using (∁?)
46+
import Data.Nat.GeneralisedArithmetic as ℕ
4647

4748

4849
open ≡-Reasoning
@@ -118,10 +119,6 @@ map-injective finj {x ∷ xs} {y ∷ ys} eq =
118119
let fx≡fy , fxs≡fys = ∷-injective eq in
119120
cong₂ _∷_ (finj fx≡fy) (map-injective finj fxs≡fys)
120121

121-
map-replicate : (f : A B) n x map f (replicate n x) ≡ replicate n (f x)
122-
map-replicate f zero x = refl
123-
map-replicate f (suc n) x = cong (_ ∷_) (map-replicate f n x)
124-
125122
------------------------------------------------------------------------
126123
-- mapMaybe
127124

@@ -621,13 +618,6 @@ sum-++ (x ∷ xs) ys = begin
621618
∈⇒∣product {n} {n ∷ ns} (here refl) = divides (product ns) (*-comm n (product ns))
622619
∈⇒∣product {n} {m ∷ ns} (there n∈ns) = ∣n⇒∣m*n m (∈⇒∣product n∈ns)
623620

624-
------------------------------------------------------------------------
625-
-- replicate
626-
627-
length-replicate : n {x : A} length (replicate n x) ≡ n
628-
length-replicate zero = refl
629-
length-replicate (suc n) = cong suc (length-replicate n)
630-
631621
------------------------------------------------------------------------
632622
-- scanr
633623

@@ -858,6 +848,48 @@ drop-drop zero n xs = refl
858848
drop-drop (suc m) n [] = drop-[] n
859849
drop-drop (suc m) n (x ∷ xs) = drop-drop m n xs
860850

851+
drop-all : (n : ℕ) (xs : List A) n ≥ length xs drop n xs ≡ []
852+
drop-all n [] _ = drop-[] n
853+
drop-all (suc n) (x ∷ xs) p = drop-all n xs (≤-pred p)
854+
855+
------------------------------------------------------------------------
856+
-- replicate
857+
858+
length-replicate : n {x : A} length (replicate n x) ≡ n
859+
length-replicate zero = refl
860+
length-replicate (suc n) = cong suc (length-replicate n)
861+
862+
lookup-replicate : n (x : A) (i : Fin n)
863+
lookup (replicate n x) (cast (sym (length-replicate n)) i) ≡ x
864+
lookup-replicate (suc n) x zero = refl
865+
lookup-replicate (suc n) x (suc i) = lookup-replicate n x i
866+
867+
map-replicate : (f : A B) n (x : A)
868+
map f (replicate n x) ≡ replicate n (f x)
869+
map-replicate f zero x = refl
870+
map-replicate f (suc n) x = cong (_ ∷_) (map-replicate f n x)
871+
872+
zipWith-replicate : n (_⊕_ : A B C) (x : A) (y : B)
873+
zipWith _⊕_ (replicate n x) (replicate n y) ≡ replicate n (x ⊕ y)
874+
zipWith-replicate zero _⊕_ x y = refl
875+
zipWith-replicate (suc n) _⊕_ x y = cong (x ⊕ y ∷_) (zipWith-replicate n _⊕_ x y)
876+
877+
------------------------------------------------------------------------
878+
-- iterate
879+
880+
length-iterate : f (x : A) n length (iterate f x n) ≡ n
881+
length-iterate f x zero = refl
882+
length-iterate f x (suc n) = cong suc (length-iterate f (f x) n)
883+
884+
iterate-id : (x : A) n iterate id x n ≡ replicate n x
885+
iterate-id x zero = refl
886+
iterate-id x (suc n) = cong (_ ∷_) (iterate-id x n)
887+
888+
lookup-iterate : f (x : A) n (i : Fin n)
889+
lookup (iterate f x n) (cast (sym (length-iterate f x n)) i) ≡ ℕ.iterate f x (toℕ i)
890+
lookup-iterate f x (suc n) zero = refl
891+
lookup-iterate f x (suc n) (suc i) = lookup-iterate f (f x) n i
892+
861893
------------------------------------------------------------------------
862894
-- splitAt
863895

src/Data/Vec/Base.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -60,9 +60,9 @@ lookup : Vec A n → Fin n → A
6060
lookup (x ∷ xs) zero = x
6161
lookup (x ∷ xs) (suc i) = lookup xs i
6262

63-
iterate : (A A) A {n} Vec A n
64-
iterate s z {zero} = []
65-
iterate s z {suc n} = z ∷ iterate s (s z)
63+
iterate : (A A) A n Vec A n
64+
iterate s z zero = []
65+
iterate s z (suc n) = z ∷ iterate s (s z) n
6666

6767
insertAt : Vec A n Fin (suc n) A Vec A (suc n)
6868
insertAt xs zero v = v ∷ xs

src/Data/Vec/Properties.agda

+24
Original file line numberDiff line numberDiff line change
@@ -31,6 +31,7 @@ open import Relation.Binary.PropositionalEquality
3131
open import Relation.Unary using (Pred; Decidable)
3232
open import Relation.Nullary.Decidable.Core using (Dec; does; yes; no; _×-dec_; map′)
3333
open import Relation.Nullary.Negation.Core using (contradiction)
34+
import Data.Nat.GeneralisedArithmetic as ℕ
3435

3536
open ≡-Reasoning
3637

@@ -1098,6 +1099,29 @@ toList-replicate : ∀ (n : ℕ) (x : A) →
10981099
toList-replicate zero x = refl
10991100
toList-replicate (suc n) x = cong (_ List.∷_) (toList-replicate n x)
11001101

1102+
------------------------------------------------------------------------
1103+
-- iterate
1104+
1105+
iterate-id : (x : A) n iterate id x n ≡ replicate x
1106+
iterate-id x zero = refl
1107+
iterate-id x (suc n) = cong (_ ∷_) (iterate-id (id x) n)
1108+
1109+
take-iterate : n f (x : A) take n (iterate f x (n + m)) ≡ iterate f x n
1110+
take-iterate zero f x = refl
1111+
take-iterate (suc n) f x = cong (_ ∷_) (take-iterate n f (f x))
1112+
1113+
drop-iterate : n f (x : A) drop n (iterate f x (n + zero)) ≡ []
1114+
drop-iterate zero f x = refl
1115+
drop-iterate (suc n) f x = drop-iterate n f (f x)
1116+
1117+
lookup-iterate : f (x : A) (i : Fin n) lookup (iterate f x n) i ≡ ℕ.iterate f x (toℕ i)
1118+
lookup-iterate f x zero = refl
1119+
lookup-iterate f x (suc i) = lookup-iterate f (f x) i
1120+
1121+
toList-iterate : f (x : A) n toList (iterate f x n) ≡ List.iterate f x n
1122+
toList-iterate f x zero = refl
1123+
toList-iterate f x (suc n) = cong (_ List.∷_) (toList-iterate f (f x) n)
1124+
11011125
------------------------------------------------------------------------
11021126
-- tabulate
11031127

0 commit comments

Comments
 (0)