@@ -43,6 +43,7 @@ open import Relation.Nullary using (¬_; Dec; does; _because_; yes; no; contradi
43
43
open import Relation.Nullary.Decidable as Decidable using (isYes; map′; ⌊_⌋; ¬?; _×-dec_)
44
44
open import Relation.Unary using (Pred; Decidable; ∁)
45
45
open import Relation.Unary.Properties using (∁?)
46
+ import Data.Nat.GeneralisedArithmetic as ℕ
46
47
47
48
48
49
open ≡-Reasoning
@@ -118,10 +119,6 @@ map-injective finj {x ∷ xs} {y ∷ ys} eq =
118
119
let fx≡fy , fxs≡fys = ∷-injective eq in
119
120
cong₂ _∷_ (finj fx≡fy) (map-injective finj fxs≡fys)
120
121
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
-
125
122
------------------------------------------------------------------------
126
123
-- mapMaybe
127
124
@@ -621,13 +618,6 @@ sum-++ (x ∷ xs) ys = begin
621
618
∈⇒∣product {n} {n ∷ ns} (here refl) = divides (product ns) (*-comm n (product ns))
622
619
∈⇒∣product {n} {m ∷ ns} (there n∈ns) = ∣n⇒∣m*n m (∈⇒∣product n∈ns)
623
620
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
-
631
621
------------------------------------------------------------------------
632
622
-- scanr
633
623
@@ -858,6 +848,48 @@ drop-drop zero n xs = refl
858
848
drop-drop (suc m) n [] = drop-[] n
859
849
drop-drop (suc m) n (x ∷ xs) = drop-drop m n xs
860
850
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
+
861
893
------------------------------------------------------------------------
862
894
-- splitAt
863
895
0 commit comments