-
Notifications
You must be signed in to change notification settings - Fork 237
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
Refactor
Data.List.Base.scan*
and their properties (#2395)
* refactor `scanr` etc. * restore `inits` etc. but lazier; plus knock-on consequences * more refactoring; plus knock-on consequences * tidy up * refactored into `Base` * ... and `Properties` * fix-up `inits` and `tails` * fix up `import`s * knock-ons * Andreas's suggestions * further, better, refactoring is possible * yet further, better, refactoring is possible: remove explicit equational reasoning! * Update CHANGELOG.md Fix deprecated names * Update Base.agda Fix deprecations * Update Properties.agda Fix deprecations * Update CHANGELOG.md Fix deprecated names
- Loading branch information
1 parent
d5407cc
commit d3f01fe
Showing
11 changed files
with
245 additions
and
49 deletions.
There are no files selected for viewing
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,49 @@ | ||
------------------------------------------------------------------------ | ||
-- The Agda standard library | ||
-- | ||
-- List scans: definitions | ||
------------------------------------------------------------------------ | ||
|
||
{-# OPTIONS --cubical-compatible --safe #-} | ||
|
||
module Data.List.Scans.Base where | ||
|
||
open import Data.List.Base as List using (List; []; _∷_) | ||
open import Data.List.NonEmpty.Base as List⁺ using (List⁺; _∷_; toList) | ||
open import Function.Base using (_∘_) | ||
open import Level using (Level) | ||
|
||
private | ||
variable | ||
a b : Level | ||
A : Set a | ||
B : Set b | ||
|
||
|
||
------------------------------------------------------------------------ | ||
-- Definitions | ||
|
||
-- Scanr | ||
|
||
module _ (f : A → B → B) where | ||
|
||
scanr⁺ : (e : B) → List A → List⁺ B | ||
scanr⁺ e [] = e ∷ [] | ||
scanr⁺ e (x ∷ xs) = let y ∷ ys = scanr⁺ e xs in f x y ∷ y ∷ ys | ||
|
||
scanr : (e : B) → List A → List B | ||
scanr e = toList ∘ scanr⁺ e | ||
|
||
-- Scanl | ||
|
||
module _ (f : A → B → A) where | ||
|
||
scanl⁺ : A → List B → List⁺ A | ||
scanl⁺ e xs = e ∷ go e xs | ||
where | ||
go : A → List B → List A | ||
go _ [] = [] | ||
go e (x ∷ xs) = let fex = f e x in fex ∷ go fex xs | ||
|
||
scanl : A → List B → List A | ||
scanl e = toList ∘ scanl⁺ e |
Oops, something went wrong.