|
8 | 8 | -- equalities than _≡_. |
9 | 9 |
|
10 | 10 | {-# OPTIONS --cubical-compatible --safe #-} |
| 11 | +{-# OPTIONS --warn=noUserWarning #-} -- for deprecated scans |
11 | 12 |
|
12 | 13 | module Data.List.Properties where |
13 | 14 |
|
@@ -809,34 +810,6 @@ sum-++ (x ∷ xs) ys = begin |
809 | 810 | ∈⇒∣product {n} {n ∷ ns} (here refl) = divides (product ns) (*-comm n (product ns)) |
810 | 811 | ∈⇒∣product {n} {m ∷ ns} (there n∈ns) = ∣n⇒∣m*n m (∈⇒∣product n∈ns) |
811 | 812 |
|
812 | | ------------------------------------------------------------------------- |
813 | | --- scanr |
814 | | - |
815 | | -scanr-defn : ∀ (f : A → B → B) (e : B) → |
816 | | - scanr f e ≗ map (foldr f e) ∘ tails |
817 | | -scanr-defn f e [] = refl |
818 | | -scanr-defn f e (x ∷ []) = refl |
819 | | -scanr-defn f e (x ∷ y∷xs@(_ ∷ _)) |
820 | | - with eq ← scanr-defn f e y∷xs |
821 | | - with z ∷ zs ← scanr f e y∷xs |
822 | | - = let z≡fy⦇f⦈xs , _ = ∷-injective eq in cong₂ (λ z → f x z ∷_) z≡fy⦇f⦈xs eq |
823 | | - |
824 | | ------------------------------------------------------------------------- |
825 | | --- scanl |
826 | | - |
827 | | -scanl-defn : ∀ (f : A → B → A) (e : A) → |
828 | | - scanl f e ≗ map (foldl f e) ∘ inits |
829 | | -scanl-defn f e [] = refl |
830 | | -scanl-defn f e (x ∷ xs) = cong (e ∷_) (begin |
831 | | - scanl f (f e x) xs |
832 | | - ≡⟨ scanl-defn f (f e x) xs ⟩ |
833 | | - map (foldl f (f e x)) (inits xs) |
834 | | - ≡⟨ refl ⟩ |
835 | | - map (foldl f e ∘ (x ∷_)) (inits xs) |
836 | | - ≡⟨ map-∘ (inits xs) ⟩ |
837 | | - map (foldl f e) (map (x ∷_) (inits xs)) |
838 | | - ∎) |
839 | | - |
840 | 813 | ------------------------------------------------------------------------ |
841 | 814 | -- applyUpTo |
842 | 815 |
|
@@ -1582,3 +1555,35 @@ map-─ = map-removeAt |
1582 | 1555 | "Warning: map-─ was deprecated in v2.0. |
1583 | 1556 | Please use map-removeAt instead." |
1584 | 1557 | #-} |
| 1558 | + |
| 1559 | +-- Version 2.1 |
| 1560 | + |
| 1561 | +scanr-defn : ∀ (f : A → B → B) (e : B) → |
| 1562 | + scanr f e ≗ map (foldr f e) ∘ tails |
| 1563 | +scanr-defn f e [] = refl |
| 1564 | +scanr-defn f e (x ∷ []) = refl |
| 1565 | +scanr-defn f e (x ∷ xs@(_ ∷ _)) |
| 1566 | + with eq ← scanr-defn f e xs |
| 1567 | + with ys@(_ ∷ _) ← scanr f e xs |
| 1568 | + = cong₂ (λ z → f x z ∷_) (∷-injectiveˡ eq) eq |
| 1569 | +{-# WARNING_ON_USAGE scanr-defn |
| 1570 | +"Warning: scanr-defn was deprecated in v2.1. |
| 1571 | +Please use Data.List.Scans.Properties.scanr-defn instead." |
| 1572 | +#-} |
| 1573 | + |
| 1574 | +scanl-defn : ∀ (f : A → B → A) (e : A) → |
| 1575 | + scanl f e ≗ map (foldl f e) ∘ inits |
| 1576 | +scanl-defn f e [] = refl |
| 1577 | +scanl-defn f e (x ∷ xs) = cong (e ∷_) (begin |
| 1578 | + scanl f (f e x) xs |
| 1579 | + ≡⟨ scanl-defn f (f e x) xs ⟩ |
| 1580 | + map (foldl f (f e x)) (inits xs) |
| 1581 | + ≡⟨ refl ⟩ |
| 1582 | + map (foldl f e ∘ (x ∷_)) (inits xs) |
| 1583 | + ≡⟨ map-∘ (inits xs) ⟩ |
| 1584 | + map (foldl f e) (map (x ∷_) (inits xs)) |
| 1585 | + ∎) |
| 1586 | +{-# WARNING_ON_USAGE scanl-defn |
| 1587 | +"Warning: scanl-defn was deprecated in v2.1. |
| 1588 | +Please use Data.List.Scans.Properties.scanl-defn instead." |
| 1589 | +#-} |
0 commit comments