Skip to content

Commit

Permalink
Merge pull request #812 from travisbrown/topic/applicative-error
Browse files Browse the repository at this point in the history
Add ApplicativeError
  • Loading branch information
adelbertc committed Jan 18, 2016
2 parents d6dce55 + 312afa7 commit 15a5e8d
Show file tree
Hide file tree
Showing 8 changed files with 204 additions and 122 deletions.
81 changes: 81 additions & 0 deletions core/src/main/scala/cats/ApplicativeError.scala
Original file line number Diff line number Diff line change
@@ -0,0 +1,81 @@
package cats

import cats.data.{Xor, XorT}

/**
* An applicative that also allows you to raise and or handle an error value.
*
* This type class allows one to abstract over error-handling applicatives.
*/
trait ApplicativeError[F[_], E] extends Applicative[F] {
/**
* Lift an error into the `F` context.
*/
def raiseError[A](e: E): F[A]

/**
* Handle any error, potentially recovering from it, by mapping it to an
* `F[A]` value.
*
* @see [[handleError]] to handle any error by simply mapping it to an `A`
* value instead of an `F[A]`.
*
* @see [[recoverWith]] to recover from only certain errors.
*/
def handleErrorWith[A](fa: F[A])(f: E => F[A]): F[A]

/**
* Handle any error, by mapping it to an `A` value.
*
* @see [[handleErrorWith]] to map to an `F[A]` value instead of simply an
* `A` value.
*
* @see [[recover]] to only recover from certain errors.
*/
def handleError[A](fa: F[A])(f: E => A): F[A] = handleErrorWith(fa)(f andThen pure)

/**
* Handle errors by turning them into [[cats.data.Xor.Left]] values.
*
* If there is no error, then an [[cats.data.Xor.Right]] value will be returned instead.
*
* All non-fatal errors should be handled by this method.
*/
def attempt[A](fa: F[A]): F[E Xor A] = handleErrorWith(
map(fa)(Xor.right[E, A])
)(e => pure(Xor.left(e)))

/**
* Similar to [[attempt]], but wraps the result in a [[cats.data.XorT]] for
* convenience.
*/
def attemptT[A](fa: F[A]): XorT[F, E, A] = XorT(attempt(fa))

/**
* Recover from certain errors by mapping them to an `A` value.
*
* @see [[handleError]] to handle any/all errors.
*
* @see [[recoverWith]] to recover from certain errors by mapping them to
* `F[A]` values.
*/
def recover[A](fa: F[A])(pf: PartialFunction[E, A]): F[A] =
handleErrorWith(fa)(e =>
(pf andThen pure) applyOrElse(e, raiseError))

/**
* Recover from certain errors by mapping them to an `F[A]` value.
*
* @see [[handleErrorWith]] to handle any/all errors.
*
* @see [[recover]] to recover from certain errors by mapping them to `A`
* values.
*/
def recoverWith[A](fa: F[A])(pf: PartialFunction[E, F[A]]): F[A] =
handleErrorWith(fa)(e =>
pf applyOrElse(e, raiseError))
}

object ApplicativeError {
def apply[F[_], E](implicit F: ApplicativeError[F, E]): ApplicativeError[F, E] = F
}
69 changes: 1 addition & 68 deletions core/src/main/scala/cats/MonadError.scala
Original file line number Diff line number Diff line change
Expand Up @@ -7,74 +7,7 @@ import cats.data.{Xor, XorT}
*
* This type class allows one to abstract over error-handling monads.
*/
trait MonadError[F[_], E] extends Monad[F] {
/**
* Lift an error into the `F` context.
*/
def raiseError[A](e: E): F[A]

/**
* Handle any error, potentially recovering from it, by mapping it to an
* `F[A]` value.
*
* @see [[handleError]] to handle any error by simply mapping it to an `A`
* value instead of an `F[A]`.
*
* @see [[recoverWith]] to recover from only certain errors.
*/
def handleErrorWith[A](fa: F[A])(f: E => F[A]): F[A]

/**
* Handle any error, by mapping it to an `A` value.
*
* @see [[handleErrorWith]] to map to an `F[A]` value instead of simply an
* `A` value.
*
* @see [[recover]] to only recover from certain errors.
*/
def handleError[A](fa: F[A])(f: E => A): F[A] = handleErrorWith(fa)(f andThen pure)

/**
* Handle errors by turning them into [[cats.data.Xor.Left]] values.
*
* If there is no error, then an [[cats.data.Xor.Right]] value will be returned instead.
*
* All non-fatal errors should be handled by this method.
*/
def attempt[A](fa: F[A]): F[E Xor A] = handleErrorWith(
map(fa)(Xor.right[E, A])
)(e => pure(Xor.left(e)))

/**
* Similar to [[attempt]], but wraps the result in a [[cats.data.XorT]] for
* convenience.
*/
def attemptT[A](fa: F[A]): XorT[F, E, A] = XorT(attempt(fa))

/**
* Recover from certain errors by mapping them to an `A` value.
*
* @see [[handleError]] to handle any/all errors.
*
* @see [[recoverWith]] to recover from certain errors by mapping them to
* `F[A]` values.
*/
def recover[A](fa: F[A])(pf: PartialFunction[E, A]): F[A] =
handleErrorWith(fa)(e =>
(pf andThen pure) applyOrElse(e, raiseError))

/**
* Recover from certain errors by mapping them to an `F[A]` value.
*
* @see [[handleErrorWith]] to handle any/all errors.
*
* @see [[recover]] to recover from certain errors by mapping them to `A`
* values.
*/
def recoverWith[A](fa: F[A])(pf: PartialFunction[E, F[A]]): F[A] =
handleErrorWith(fa)(e =>
pf applyOrElse(e, raiseError))
}
trait MonadError[F[_], E] extends ApplicativeError[F, E] with Monad[F]

object MonadError {
def apply[F[_], E](implicit F: MonadError[F, E]): MonadError[F, E] = F
Expand Down
11 changes: 9 additions & 2 deletions core/src/main/scala/cats/data/Validated.scala
Original file line number Diff line number Diff line change
Expand Up @@ -234,8 +234,8 @@ private[data] sealed abstract class ValidatedInstances extends ValidatedInstance
override def leftMap[A, B, C](fab: Validated[A, B])(f: A => C): Validated[C, B] = fab.leftMap(f)
}

implicit def validatedInstances[E](implicit E: Semigroup[E]): Traverse[Validated[E, ?]] with Applicative[Validated[E, ?]] =
new Traverse[Validated[E, ?]] with Applicative[Validated[E,?]] {
implicit def validatedInstances[E](implicit E: Semigroup[E]): Traverse[Validated[E, ?]] with ApplicativeError[Validated[E, ?], E] =
new Traverse[Validated[E, ?]] with ApplicativeError[Validated[E, ?], E] {
def traverse[F[_]: Applicative, A, B](fa: Validated[E,A])(f: A => F[B]): F[Validated[E,B]] =
fa.traverse(f)

Expand All @@ -256,6 +256,13 @@ private[data] sealed abstract class ValidatedInstances extends ValidatedInstance

def product[A, B](fa: Validated[E, A], fb: Validated[E, B]): Validated[E, (A, B)] =
fa.product(fb)(E)

def handleErrorWith[A](fa: Validated[E, A])(f: E => Validated[E, A]): Validated[E, A] =
fa match {
case Validated.Invalid(e) => f(e)
case v @ Validated.Valid(_) => v
}
def raiseError[A](e: E): Validated[E, A] = Validated.Invalid(e)
}
}

Expand Down
44 changes: 44 additions & 0 deletions laws/src/main/scala/cats/laws/ApplicativeErrorLaws.scala
Original file line number Diff line number Diff line change
@@ -0,0 +1,44 @@
package cats
package laws

import cats.data.{Xor, XorT}

// Taken from http://functorial.com/psc-pages/docs/Control/Monad/Error/Class/index.html
trait ApplicativeErrorLaws[F[_], E] extends ApplicativeLaws[F] {
implicit override def F: ApplicativeError[F, E]

def applicativeErrorHandleWith[A](e: E, f: E => F[A]): IsEq[F[A]] =
F.handleErrorWith(F.raiseError[A](e))(f) <-> f(e)

def applicativeErrorHandle[A](e: E, f: E => A): IsEq[F[A]] =
F.handleError(F.raiseError[A](e))(f) <-> F.pure(f(e))

def handleErrorWithPure[A](a: A, f: E => F[A]): IsEq[F[A]] =
F.handleErrorWith(F.pure(a))(f) <-> F.pure(a)

def handleErrorPure[A](a: A, f: E => A): IsEq[F[A]] =
F.handleError(F.pure(a))(f) <-> F.pure(a)

def raiseErrorAttempt(e: E): IsEq[F[E Xor Unit]] =
F.attempt(F.raiseError[Unit](e)) <-> F.pure(Xor.left(e))

def pureAttempt[A](a: A): IsEq[F[E Xor A]] =
F.attempt(F.pure(a)) <-> F.pure(Xor.right(a))

def handleErrorWithConsistentWithRecoverWith[A](fa: F[A], f: E => F[A]): IsEq[F[A]] =
F.handleErrorWith(fa)(f) <-> F.recoverWith(fa)(PartialFunction(f))

def handleErrorConsistentWithRecover[A](fa: F[A], f: E => A): IsEq[F[A]] =
F.handleError(fa)(f) <-> F.recover(fa)(PartialFunction(f))

def recoverConsistentWithRecoverWith[A](fa: F[A], pf: PartialFunction[E, A]): IsEq[F[A]] =
F.recover(fa)(pf) <-> F.recoverWith(fa)(pf andThen F.pure)

def attemptConsistentWithAttemptT[A](fa: F[A]): IsEq[XorT[F, E, A]] =
XorT(F.attempt(fa)) <-> F.attemptT(fa)
}

object ApplicativeErrorLaws {
def apply[F[_], E](implicit ev: ApplicativeError[F, E]): ApplicativeErrorLaws[F, E] =
new ApplicativeErrorLaws[F, E] { def F: ApplicativeError[F, E] = ev }
}
34 changes: 1 addition & 33 deletions laws/src/main/scala/cats/laws/MonadErrorLaws.scala
Original file line number Diff line number Diff line change
@@ -1,44 +1,12 @@
package cats
package laws

import cats.data.{Xor, XorT}

// Taken from http://functorial.com/psc-pages/docs/Control/Monad/Error/Class/index.html
trait MonadErrorLaws[F[_], E] extends MonadLaws[F] {
trait MonadErrorLaws[F[_], E] extends ApplicativeErrorLaws[F, E] with MonadLaws[F] {
implicit override def F: MonadError[F, E]

def monadErrorLeftZero[A, B](e: E, f: A => F[B]): IsEq[F[B]] =
F.flatMap(F.raiseError[A](e))(f) <-> F.raiseError[B](e)

def monadErrorHandleWith[A](e: E, f: E => F[A]): IsEq[F[A]] =
F.handleErrorWith(F.raiseError[A](e))(f) <-> f(e)

def monadErrorHandle[A](e: E, f: E => A): IsEq[F[A]] =
F.handleError(F.raiseError[A](e))(f) <-> F.pure(f(e))

def handleErrorWithPure[A](a: A, f: E => F[A]): IsEq[F[A]] =
F.handleErrorWith(F.pure(a))(f) <-> F.pure(a)

def handleErrorPure[A](a: A, f: E => A): IsEq[F[A]] =
F.handleError(F.pure(a))(f) <-> F.pure(a)

def raiseErrorAttempt(e: E): IsEq[F[E Xor Unit]] =
F.attempt(F.raiseError[Unit](e)) <-> F.pure(Xor.left(e))

def pureAttempt[A](a: A): IsEq[F[E Xor A]] =
F.attempt(F.pure(a)) <-> F.pure(Xor.right(a))

def handleErrorWithConsistentWithRecoverWith[A](fa: F[A], f: E => F[A]): IsEq[F[A]] =
F.handleErrorWith(fa)(f) <-> F.recoverWith(fa)(PartialFunction(f))

def handleErrorConsistentWithRecover[A](fa: F[A], f: E => A): IsEq[F[A]] =
F.handleError(fa)(f) <-> F.recover(fa)(PartialFunction(f))

def recoverConsistentWithRecoverWith[A](fa: F[A], pf: PartialFunction[E, A]): IsEq[F[A]] =
F.recover(fa)(pf) <-> F.recoverWith(fa)(pf andThen F.pure)

def attemptConsistentWithAttemptT[A](fa: F[A]): IsEq[XorT[F, E, A]] =
XorT(F.attempt(fa)) <-> F.attemptT(fa)
}

object MonadErrorLaws {
Expand Down
Original file line number Diff line number Diff line change
@@ -0,0 +1,57 @@
package cats
package laws
package discipline

import cats.data.{ Xor, XorT }
import cats.laws.discipline.MonoidalTests.Isomorphisms
import cats.laws.discipline.arbitrary._
import cats.laws.discipline.eq.unitEq
import org.scalacheck.{Arbitrary, Prop}
import org.scalacheck.Prop.forAll

trait ApplicativeErrorTests[F[_], E] extends ApplicativeTests[F] {
def laws: ApplicativeErrorLaws[F, E]

def applicativeError[A: Arbitrary: Eq, B: Arbitrary: Eq, C: Arbitrary: Eq](implicit
ArbFA: Arbitrary[F[A]],
ArbFB: Arbitrary[F[B]],
ArbFC: Arbitrary[F[C]],
ArbFAtoB: Arbitrary[F[A => B]],
ArbFBtoC: Arbitrary[F[B => C]],
ArbE: Arbitrary[E],
EqFA: Eq[F[A]],
EqFB: Eq[F[B]],
EqFC: Eq[F[C]],
EqE: Eq[E],
EqFXorEU: Eq[F[E Xor Unit]],
EqFXorEA: Eq[F[E Xor A]],
EqXorTFEA: Eq[XorT[F, E, A]],
EqFABC: Eq[F[(A, B, C)]],
iso: Isomorphisms[F]
): RuleSet = {
new RuleSet {
def name: String = "applicativeError"
def bases: Seq[(String, RuleSet)] = Nil
def parents: Seq[RuleSet] = Seq(applicative[A, B, C])
def props: Seq[(String, Prop)] = Seq(
"applicativeError handleWith" -> forAll(laws.applicativeErrorHandleWith[A] _),
"applicativeError handle" -> forAll(laws.applicativeErrorHandle[A] _),
"applicativeError handleErrorWith pure" -> forAll(laws.handleErrorWithPure[A] _),
"applicativeError handleError pure" -> forAll(laws.handleErrorPure[A] _),
"applicativeError raiseError attempt" -> forAll(laws.raiseErrorAttempt _),
"applicativeError pure attempt" -> forAll(laws.pureAttempt[A] _),
"applicativeError handleErrorWith consistent with recoverWith" -> forAll(laws.handleErrorWithConsistentWithRecoverWith[A] _),
"applicativeError handleError consistent with recover" -> forAll(laws.handleErrorConsistentWithRecover[A] _),
"applicativeError recover consistent with recoverWith" -> forAll(laws.recoverConsistentWithRecoverWith[A] _),
"applicativeError attempt consistent with attemptT" -> forAll(laws.attemptConsistentWithAttemptT[A] _)
)
}
}
}

object ApplicativeErrorTests {
def apply[F[_], E](implicit FE: ApplicativeError[F, E]): ApplicativeErrorTests[F, E] =
new ApplicativeErrorTests[F, E] {
def laws: ApplicativeErrorLaws[F, E] = ApplicativeErrorLaws[F, E]
}
}
17 changes: 3 additions & 14 deletions laws/src/main/scala/cats/laws/discipline/MonadErrorTests.scala
Original file line number Diff line number Diff line change
Expand Up @@ -4,12 +4,11 @@ package discipline

import cats.data.{ Xor, XorT }
import cats.laws.discipline.MonoidalTests.Isomorphisms
import cats.laws.discipline.arbitrary._
import cats.laws.discipline.eq.unitEq
import org.scalacheck.{Arbitrary, Prop}
import org.scalacheck.Prop.forAll

trait MonadErrorTests[F[_], E] extends MonadTests[F] {
trait MonadErrorTests[F[_], E] extends ApplicativeErrorTests[F, E] with MonadTests[F] {
def laws: MonadErrorLaws[F, E]

def monadError[A: Arbitrary: Eq, B: Arbitrary: Eq, C: Arbitrary: Eq](implicit
Expand All @@ -32,19 +31,9 @@ trait MonadErrorTests[F[_], E] extends MonadTests[F] {
new RuleSet {
def name: String = "monadError"
def bases: Seq[(String, RuleSet)] = Nil
def parents: Seq[RuleSet] = Seq(monad[A, B, C])
def parents: Seq[RuleSet] = Seq(applicativeError[A, B, C], monad[A, B, C])
def props: Seq[(String, Prop)] = Seq(
"monadError left zero" -> forAll(laws.monadErrorLeftZero[A, B] _),
"monadError handleWith" -> forAll(laws.monadErrorHandleWith[A] _),
"monadError handle" -> forAll(laws.monadErrorHandle[A] _),
"monadError handleErrorWith pure" -> forAll(laws.handleErrorWithPure[A] _),
"monadError handleError pure" -> forAll(laws.handleErrorPure[A] _),
"monadError raiseError attempt" -> forAll(laws.raiseErrorAttempt _),
"monadError pure attempt" -> forAll(laws.pureAttempt[A] _),
"monadError handleErrorWith consistent with recoverWith" -> forAll(laws.handleErrorWithConsistentWithRecoverWith[A] _),
"monadError handleError consistent with recover" -> forAll(laws.handleErrorConsistentWithRecover[A] _),
"monadError recover consistent with recoverWith" -> forAll(laws.recoverConsistentWithRecoverWith[A] _),
"monadError attempt consistent with attemptT" -> forAll(laws.attemptConsistentWithAttemptT[A] _)
"monadError left zero" -> forAll(laws.monadErrorLeftZero[A, B] _)
)
}
}
Expand Down
13 changes: 8 additions & 5 deletions tests/src/test/scala/cats/tests/ValidatedTests.scala
Original file line number Diff line number Diff line change
@@ -1,9 +1,9 @@
package cats
package tests

import cats.data.{NonEmptyList, Validated, ValidatedNel, Xor}
import cats.data.{NonEmptyList, Validated, ValidatedNel, Xor, XorT}
import cats.data.Validated.{Valid, Invalid}
import cats.laws.discipline.{BifunctorTests, TraverseTests, ApplicativeTests, SerializableTests, MonoidalTests}
import cats.laws.discipline.{BifunctorTests, TraverseTests, ApplicativeErrorTests, SerializableTests, MonoidalTests}
import org.scalacheck.{Gen, Arbitrary}
import org.scalacheck.Arbitrary._
import cats.laws.discipline.arbitrary._
Expand All @@ -17,12 +17,15 @@ class ValidatedTests extends CatsSuite {
checkAll("Validated[String, Int]", MonoidalTests[Validated[String,?]].monoidal[Int, Int, Int])
checkAll("Monoidal[Validated[String,?]]", SerializableTests.serializable(Monoidal[Validated[String,?]]))

checkAll("Validated[String, Int]", ApplicativeTests[Validated[String,?]].applicative[Int, Int, Int])
checkAll("Validated[?, ?]", BifunctorTests[Validated].bifunctor[Int, Int, Int, Int, Int, Int])
checkAll("Applicative[Validated[String,?]]", SerializableTests.serializable(Applicative[Validated[String,?]]))

implicit val eq0 = XorT.xorTEq[Validated[String, ?], String, Int]

checkAll("Validated[String, Int]", ApplicativeErrorTests[Validated[String, ?], String].applicativeError[Int, Int, Int])
checkAll("ApplicativeError[Xor, String]", SerializableTests.serializable(ApplicativeError[Validated[String, ?], String]))

checkAll("Validated[String, Int] with Option", TraverseTests[Validated[String,?]].traverse[Int, Int, Int, Int, Option, Option])
checkAll("Traverse[Validated[String,?]]", SerializableTests.serializable(Traverse[Validated[String,?]]))
checkAll("Traverse[Validated[String, ?]]", SerializableTests.serializable(Traverse[Validated[String,?]]))

checkAll("Validated[String, Int]", OrderLaws[Validated[String, Int]].order)
checkAll("Order[Validated[String, Int]]", SerializableTests.serializable(Order[Validated[String, Int]]))
Expand Down

0 comments on commit 15a5e8d

Please sign in to comment.