Skip to content

Commit

Permalink
Define Apply.product in term of ap and map
Browse files Browse the repository at this point in the history
- Remove all equivalent definition of `product` introduced by typelevel#555
- Changes `cats.syntax.all._` to `import cats.syntax.monoidal._` in doc
- Minor type argument spacing
  • Loading branch information
OlivierBlanvillain committed Jan 9, 2016
1 parent 0529841 commit 05f074f
Show file tree
Hide file tree
Showing 13 changed files with 43 additions and 84 deletions.
10 changes: 4 additions & 6 deletions core/src/main/scala/cats/Apply.scala
Expand Up @@ -20,15 +20,15 @@ trait Apply[F[_]] extends Functor[F] with Monoidal[F] with ApplyArityFunctions[F
* ap2 is a binary version of ap, defined in terms of ap.
*/
def ap2[A, B, Z](fa: F[A], fb: F[B])(ff: F[(A, B) => Z]): F[Z] =
map(product(fa, product(fb, ff))) { case (a, (b, f)) => f(a, b) }
ap(fb)(ap(fa)(map(ff)(ff => (a: A) => (b: B) => ff(a, b))))

/**
* Applies the pure (binary) function f to the effectful values fa and fb.
*
* map2 can be seen as a binary version of [[cats.Functor]]#map.
*/
def map2[A, B, Z](fa: F[A], fb: F[B])(f: (A, B) => Z): F[Z] =
map(product(fa, fb)) { case (a, b) => f(a, b) }
ap(fb)(map(fa)(a => (b: B) => f(a, b)))

/**
* Two sequentially dependent Applys can be composed.
Expand All @@ -46,6 +46,8 @@ trait Apply[F[_]] extends Functor[F] with Monoidal[F] with ApplyArityFunctions[F
def G: Apply[G] = GG
}

override def product[A, B](fa: F[A], fb: F[B]): F[(A, B)] =
ap(fa)(map(fb)(b => (a => (a, b))))
}

trait CompositeApply[F[_], G[_]]
Expand All @@ -55,8 +57,4 @@ trait CompositeApply[F[_], G[_]]

def ap[A, B](fa: F[G[A]])(f: F[G[A => B]]): F[G[B]] =
F.ap(fa)(F.map(f)(gab => G.ap(_)(gab)))

def product[A, B](fa: F[G[A]], fb: F[G[B]]): F[G[(A, B)]] =
F.map2(fa, fb)(G.product)

}
21 changes: 10 additions & 11 deletions core/src/main/scala/cats/FlatMap.scala
Expand Up @@ -3,17 +3,17 @@ package cats
import simulacrum.typeclass

/**
* FlatMap type class gives us flatMap, which allows us to have a value
* in a context (F[A]) and then feed that into a function that takes
* a normal value and returns a value in a context (A => F[B]).
* FlatMap type class gives us flatMap, which allows us to have a value
* in a context (F[A]) and then feed that into a function that takes
* a normal value and returns a value in a context (A => F[B]).
*
* One motivation for separating this out from Monad is that there are
* situations where we can implement flatMap but not pure. For example,
* we can implement map or flatMap that transforms the values of Map[K, ?],
* but we can't implement pure (because we wouldn't know what key to use
* when instantiating the new Map).
* One motivation for separating this out from Monad is that there are
* situations where we can implement flatMap but not pure. For example,
* we can implement map or flatMap that transforms the values of Map[K, ?],
* but we can't implement pure (because we wouldn't know what key to use
* when instantiating the new Map).
*
* @see See [[https://github.com/non/cats/issues/3]] for some discussion.
* @see See [[https://github.com/non/cats/issues/3]] for some discussion.
*
* Must obey the laws defined in cats.laws.FlatMapLaws.
*/
Expand Down Expand Up @@ -41,7 +41,6 @@ import simulacrum.typeclass
/**
* `if` lifted into monad.
*/
def ifM[B](fa: F[Boolean])(ifTrue: => F[B], ifFalse: => F[B]): F[B] = {
def ifM[B](fa: F[Boolean])(ifTrue: => F[B], ifFalse: => F[B]): F[B] =
flatMap(fa)(if (_) ifTrue else ifFalse)
}
}
6 changes: 0 additions & 6 deletions core/src/main/scala/cats/data/Const.scala
Expand Up @@ -91,9 +91,6 @@ private[data] sealed abstract class ConstInstances0 extends ConstInstances1 {

def map[A, B](fa: Const[C, A])(f: A => B): Const[C, B] =
fa.retag[B]

def product[A, B](fa: Const[C, A], fb: Const[C, B]): Const[C, (A, B)] =
fa.retag[(A, B)] combine fb.retag[(A, B)]
}
}

Expand All @@ -107,9 +104,6 @@ private[data] sealed abstract class ConstInstances1 {
def ap[A, B](fa: Const[C, A])(f: Const[C, A => B]): Const[C, B] =
fa.retag[B] combine f.retag[B]

def product[A, B](fa: Const[C, A], fb: Const[C, B]): Const[C, (A, B)] =
fa.retag[(A, B)] combine fb.retag[(A, B)]

def map[A, B](fa: Const[C, A])(f: A => B): Const[C, B] =
fa.retag[B]
}
Expand Down
4 changes: 0 additions & 4 deletions core/src/main/scala/cats/data/Func.scala
Expand Up @@ -62,8 +62,6 @@ sealed trait FuncApply[F[_], C] extends Apply[Lambda[X => Func[F, C, X]]] with F
def F: Apply[F]
def ap[A, B](fa: Func[F, C, A])(f: Func[F, C, A => B]): Func[F, C, B] =
Func.func(c => F.ap(fa.run(c))(f.run(c)))
def product[A, B](fa: Func[F, C, A], fb: Func[F, C, B]): Func[F, C, (A, B)] =
Func.func(c => F.product(fa.run(c), fb.run(c)))
}

sealed trait FuncApplicative[F[_], C] extends Applicative[Lambda[X => Func[F, C, X]]] with FuncApply[F, C] {
Expand Down Expand Up @@ -125,8 +123,6 @@ private[data] sealed trait AppFuncApplicative[F[_], C] extends Applicative[Lambd
fa.map(f)
def ap[A, B](fa: AppFunc[F, C, A])(f: AppFunc[F, C, A => B]): AppFunc[F, C, B] =
Func.appFunc[F, C, B](c => F.ap(fa.run(c))(f.run(c)))(F)
def product[A, B](fa: AppFunc[F, C, A], fb: AppFunc[F, C, B]): AppFunc[F, C, (A, B)] =
Func.appFunc[F, C, (A, B)](c => F.product(fa.run(c), fb.run(c)))(F)
def pure[A](a: A): AppFunc[F, C, A] =
Func.appFunc[F, C, A](c => F.pure(a))(F)
}
6 changes: 0 additions & 6 deletions core/src/main/scala/cats/data/Kleisli.scala
Expand Up @@ -148,9 +148,6 @@ private[data] sealed abstract class KleisliInstances1 extends KleisliInstances2

def map[B, C](fb: Kleisli[F, A, B])(f: B => C): Kleisli[F, A, C] =
fb.map(f)

def product[B, C](fb: Kleisli[F, A, B], fc: Kleisli[F, A, C]): Kleisli[F, A, (B, C)] =
Kleisli(a => Applicative[F].product(fb.run(a), fc.run(a)))
}
}

Expand All @@ -159,9 +156,6 @@ private[data] sealed abstract class KleisliInstances2 extends KleisliInstances3
def ap[B, C](fa: Kleisli[F, A, B])(f: Kleisli[F, A, B => C]): Kleisli[F, A, C] =
fa(f)

def product[B, C](fb: Kleisli[F, A, B], fc: Kleisli[F, A, C]): Kleisli[F, A, (B, C)] =
Kleisli(a => Apply[F].product(fb.run(a), fc.run(a)))

def map[B, C](fa: Kleisli[F, A, B])(f: B => C): Kleisli[F, A, C] =
fa.map(f)
}
Expand Down
2 changes: 0 additions & 2 deletions core/src/main/scala/cats/data/Prod.scala
Expand Up @@ -79,8 +79,6 @@ sealed trait ProdApply[F[_], G[_]] extends Apply[Lambda[X => Prod[F, G, X]]] wit
def G: Apply[G]
def ap[A, B](fa: Prod[F, G, A])(f: Prod[F, G, A => B]): Prod[F, G, B] =
Prod(F.ap(fa.first)(f.first), G.ap(fa.second)(f.second))
def product[A, B](fa: Prod[F, G, A], fb: Prod[F, G, B]): Prod[F, G, (A, B)] =
Prod(F.product(fa.first, fb.first), G.product(fa.second, fb.second))
}

sealed trait ProdApplicative[F[_], G[_]] extends Applicative[Lambda[X => Prod[F, G, X]]] with ProdApply[F, G] {
Expand Down
54 changes: 27 additions & 27 deletions core/src/main/scala/cats/data/Validated.scala
Expand Up @@ -79,7 +79,7 @@ sealed abstract class Validated[+E, +A] extends Product with Serializable {
* Convert to an Xor, apply a function, convert back. This is handy
* when you want to use the Monadic properties of the Xor type.
*/
def withXor[EE,B](f: (E Xor A) => (EE Xor B)): Validated[EE,B] =
def withXor[EE,B](f: (E Xor A) => (EE Xor B)): Validated[EE, B] =
f(toXor).toValidated

/**
Expand Down Expand Up @@ -109,7 +109,7 @@ sealed abstract class Validated[+E, +A] extends Product with Serializable {
* From Apply:
* if both the function and this value are Valid, apply the function
*/
def ap[EE >: E, B](f: Validated[EE, A => B])(implicit EE: Semigroup[EE]): Validated[EE,B] =
def ap[EE >: E, B](f: Validated[EE, A => B])(implicit EE: Semigroup[EE]): Validated[EE, B] =
(this, f) match {
case (Valid(a), Valid(f)) => Valid(f(a))
case (Invalid(e1), Invalid(e2)) => Invalid(EE.combine(e2, e1))
Expand All @@ -131,20 +131,20 @@ sealed abstract class Validated[+E, +A] extends Product with Serializable {
/**
* Apply a function to a Valid value, returning a new Valid value
*/
def map[B](f: A => B): Validated[E,B] = bimap(identity, f)
def map[B](f: A => B): Validated[E, B] = bimap(identity, f)

/**
* Apply a function to an Invalid value, returning a new Invalid value.
* Or, if the original valid was Valid, return it.
*/
def leftMap[EE](f: E => EE): Validated[EE,A] = bimap(f, identity)
def leftMap[EE](f: E => EE): Validated[EE, A] = bimap(f, identity)

/**
* When Valid, apply the function, marking the result as valid
* inside the Applicative's context,
* when Invalid, lift the Error into the Applicative's context
*/
def traverse[F[_], EE >: E, B](f: A => F[B])(implicit F: Applicative[F]): F[Validated[EE,B]] =
def traverse[F[_], EE >: E, B](f: A => F[B])(implicit F: Applicative[F]): F[Validated[EE, B]] =
fold(e => F.pure(Invalid(e)),
a => F.map(f(a))(Valid.apply))

Expand Down Expand Up @@ -218,14 +218,14 @@ private[data] sealed abstract class ValidatedInstances extends ValidatedInstance
def combine(x: Validated[A, B], y: Validated[A, B]): Validated[A, B] = x combine y
}

implicit def validatedOrder[A: Order, B: Order]: Order[Validated[A,B]] = new Order[Validated[A,B]] {
def compare(x: Validated[A,B], y: Validated[A,B]): Int = x compare y
override def partialCompare(x: Validated[A,B], y: Validated[A,B]): Double = x partialCompare y
override def eqv(x: Validated[A,B], y: Validated[A,B]): Boolean = x === y
implicit def validatedOrder[A: Order, B: Order]: Order[Validated[A, B]] = new Order[Validated[A, B]] {
def compare(x: Validated[A, B], y: Validated[A, B]): Int = x compare y
override def partialCompare(x: Validated[A, B], y: Validated[A, B]): Double = x partialCompare y
override def eqv(x: Validated[A, B], y: Validated[A, B]): Boolean = x === y
}

implicit def validatedShow[A, B](implicit A: Show[A], B: Show[B]): Show[Validated[A,B]] = new Show[Validated[A,B]] {
def show(f: Validated[A,B]): String = f.show
implicit def validatedShow[A, B](implicit A: Show[A], B: Show[B]): Show[Validated[A, B]] = new Show[Validated[A, B]] {
def show(f: Validated[A, B]): String = f.show
}

implicit def validatedBifunctor: Bifunctor[Validated] =
Expand All @@ -242,7 +242,7 @@ private[data] sealed abstract class ValidatedInstances extends ValidatedInstance
def foldLeft[A, B](fa: Validated[E,A], b: B)(f: (B, A) => B): B =
fa.foldLeft(b)(f)

def foldRight[A,B](fa: Validated[E,A], lb: Eval[B])(f: (A, Eval[B]) => Eval[B]): Eval[B] =
def foldRight[A, B](fa: Validated[E,A], lb: Eval[B])(f: (A, Eval[B]) => Eval[B]): Eval[B] =
fa.foldRight(lb)(f)

def pure[A](a: A): Validated[E,A] =
Expand All @@ -251,10 +251,10 @@ private[data] sealed abstract class ValidatedInstances extends ValidatedInstance
override def map[A, B](fa: Validated[E,A])(f: A => B): Validated[E, B] =
fa.map(f)

def ap[A,B](fa: Validated[E,A])(f: Validated[E,A=>B]): Validated[E, B] =
def ap[A, B](fa: Validated[E,A])(f: Validated[E,A=>B]): Validated[E, B] =
fa.ap(f)(E)

def product[A, B](fa: Validated[E, A], fb: Validated[E, B]): Validated[E, (A, B)] =
override def product[A, B](fa: Validated[E, A], fb: Validated[E, B]): Validated[E, (A, B)] =
fa.product(fb)(E)
}
}
Expand All @@ -266,26 +266,26 @@ private[data] sealed abstract class ValidatedInstances1 extends ValidatedInstanc
def combine(x: Validated[A, B], y: Validated[A, B]): Validated[A, B] = x combine y
}

implicit def validatedPartialOrder[A: PartialOrder, B: PartialOrder]: PartialOrder[Validated[A,B]] =
new PartialOrder[Validated[A,B]] {
def partialCompare(x: Validated[A,B], y: Validated[A,B]): Double = x partialCompare y
override def eqv(x: Validated[A,B], y: Validated[A,B]): Boolean = x === y
implicit def validatedPartialOrder[A: PartialOrder, B: PartialOrder]: PartialOrder[Validated[A, B]] =
new PartialOrder[Validated[A, B]] {
def partialCompare(x: Validated[A, B], y: Validated[A, B]): Double = x partialCompare y
override def eqv(x: Validated[A, B], y: Validated[A, B]): Boolean = x === y
}
}

private[data] sealed abstract class ValidatedInstances2 {
implicit def validatedEq[A: Eq, B: Eq]: Eq[Validated[A,B]] =
new Eq[Validated[A,B]] {
def eqv(x: Validated[A,B], y: Validated[A,B]): Boolean = x === y
implicit def validatedEq[A: Eq, B: Eq]: Eq[Validated[A, B]] =
new Eq[Validated[A, B]] {
def eqv(x: Validated[A, B], y: Validated[A, B]): Boolean = x === y
}
}

trait ValidatedFunctions {
def invalid[A, B](a: A): Validated[A,B] = Validated.Invalid(a)
def invalid[A, B](a: A): Validated[A, B] = Validated.Invalid(a)

def invalidNel[A, B](a: A): ValidatedNel[A, B] = Validated.Invalid(NonEmptyList(a))

def valid[A, B](b: B): Validated[A,B] = Validated.Valid(b)
def valid[A, B](b: B): Validated[A, B] = Validated.Valid(b)

/**
* Evaluates the specified block, catching exceptions of the specified type and returning them on the invalid side of
Expand Down Expand Up @@ -325,13 +325,13 @@ trait ValidatedFunctions {
}

/**
* Converts an `Either[A, B]` to an `Validated[A,B]`.
* Converts an `Either[A, B]` to an `Validated[A, B]`.
*/
def fromEither[A, B](e: Either[A, B]): Validated[A,B] = e.fold(invalid, valid)
def fromEither[A, B](e: Either[A, B]): Validated[A, B] = e.fold(invalid, valid)

/**
* Converts an `Option[B]` to an `Validated[A,B]`, where the provided `ifNone` values is returned on
* Converts an `Option[B]` to an `Validated[A, B]`, where the provided `ifNone` values is returned on
* the invalid of the `Validated` when the specified `Option` is `None`.
*/
def fromOption[A, B](o: Option[B], ifNone: => A): Validated[A,B] = o.fold(invalid[A, B](ifNone))(valid)
def fromOption[A, B](o: Option[B], ifNone: => A): Validated[A, B] = o.fold(invalid[A, B](ifNone))(valid)
}
2 changes: 0 additions & 2 deletions core/src/main/scala/cats/data/WriterT.scala
Expand Up @@ -157,8 +157,6 @@ private[data] sealed trait WriterTApply[F[_], L] extends WriterTFunctor[F, L] wi

def ap[A, B](fa: WriterT[F, L, A])(f: WriterT[F, L, A => B]): WriterT[F, L, B] =
fa ap f
def product[A, B](fa: WriterT[F, L, A], fb: WriterT[F, L, B]): WriterT[F, L, (A, B)] =
WriterT(F0.map(F0.product(fa.run, fb.run)) { case ((l1, a), (l2, b)) => (L0.combine(l1, l2), (a, b)) })
}

private[data] sealed trait WriterTFlatMap[F[_], L] extends WriterTApply[F, L] with FlatMap[WriterT[F, L, ?]] {
Expand Down
3 changes: 1 addition & 2 deletions core/src/main/scala/cats/free/FreeApplicative.scala
Expand Up @@ -48,7 +48,7 @@ sealed abstract class FreeApplicative[F[_], A] extends Product with Serializable
}

/** Interpret this algebra into a Monoid */
final def analyze[M:Monoid](f: F ~> λ[α => M]): M =
final def analyze[M: Monoid](f: F ~> λ[α => M]): M =
foldMap[Const[M, ?]](new (F ~> Const[M, ?]) {
def apply[X](x: F[X]): Const[M,X] = Const(f(x))
}).getConst
Expand Down Expand Up @@ -79,7 +79,6 @@ object FreeApplicative {

implicit final def freeApplicative[S[_]]: Applicative[FA[S, ?]] = {
new Applicative[FA[S, ?]] {
def product[A, B](fa: FA[S, A], fb: FA[S, B]): FA[S, (A, B)] = ap(fb)(fa.map(a => b => (a, b)))
def map[A, B](fa: FA[S, A])(f: A => B): FA[S, B] = fa.map(f)
override def ap[A, B](fa: FA[S, A])(f: FA[S, A => B]): FA[S, B] = fa.ap(f)
def pure[A](a: A): FA[S, A] = Pure(a)
Expand Down
8 changes: 1 addition & 7 deletions docs/src/main/tut/apply.md
Expand Up @@ -26,19 +26,13 @@ implicit val optionApply: Apply[Option] = new Apply[Option] {
fa.flatMap (a => f.map (ff => ff(a)))
def map[A,B](fa: Option[A])(f: A => B): Option[B] = fa map f
def product[A, B](fa: Option[A], fb: Option[B]): Option[(A, B)] =
fa.flatMap(a => fb.map(b => (a, b)))
}
implicit val listApply: Apply[List] = new Apply[List] {
def ap[A, B](fa: List[A])(f: List[A => B]): List[B] =
fa.flatMap (a => f.map (ff => ff(a)))
def map[A,B](fa: List[A])(f: A => B): List[B] = fa map f
def product[A, B](fa: List[A], fb: List[B]): List[(A, B)] =
fa.zip(fb)
}
```

Expand Down Expand Up @@ -121,7 +115,7 @@ Apply[Option].tuple3(Some(1), Some(2), Some(3))

The `|@|` operator offers an alternative syntax for the higher-arity `Apply`
functions (`apN`, `mapN` and `tupleN`).
In order to use it, first import `cats.syntax.all._` or `cats.syntax.apply._`.
In order to use it, first import `import cats.syntax.monoidal._`.
Here we see that the following two functions, `f1` and `f2`, are equivalent:

```tut
Expand Down
5 changes: 0 additions & 5 deletions docs/src/main/tut/const.md
Expand Up @@ -221,8 +221,6 @@ implicit def constApplicative[Z]: Applicative[Const[Z, ?]] =
def ap[A, B](fa: Const[Z, A])(f: Const[Z, A => B]): Const[Z, B] = ???
def map[A, B](fa: Const[Z, A])(f: A => B): Const[Z, B] = ???
def product[A, B](fa: Const[Z, A],fb: Const[Z, B]): Const[Z, (A, B)] = ???
}
```

Expand All @@ -249,9 +247,6 @@ implicit def constApplicative[Z : Monoid]: Applicative[Const[Z, ?]] =
def map[A, B](fa: Const[Z, A])(f: A => B): Const[Z, B] =
Const(fa.getConst)
def product[A, B](fa: Const[Z, A],fb: Const[Z, B]): Const[Z, (A, B)] =
Const(Monoid[Z].combine(fa.getConst, fb.getConst))
}
```

Expand Down
1 change: 0 additions & 1 deletion docs/src/main/tut/validated.md
Expand Up @@ -200,7 +200,6 @@ implicit def validatedApplicative[E : Semigroup]: Applicative[Validated[E, ?]] =
def pure[A](x: A): Validated[E, A] = Validated.valid(x)
def map[A, B](fa: Validated[E, A])(f: A => B): Validated[E, B] = fa.map(f)
def product[A, B](fa: Validated[E, A], fb: Validated[E, B]): Validated[E, (A, B)] = ap(fb)(fa.map(a => b => (a, b)))
}
```

Expand Down
5 changes: 0 additions & 5 deletions laws/src/main/scala/cats/laws/TraverseLaws.scala
Expand Up @@ -49,11 +49,6 @@ trait TraverseLaws[F[_]] extends FunctorLaws[F] with FoldableLaws[F] {
val (mx, nx) = fx
(M.map(mx)(f), N.map(nx)(f))
}
def product[X, Y](fx: MN[X], fy: MN[Y]): MN[(X, Y)] = {
val (mx, nx) = fx
val (my, ny) = fy
(M.product(mx, my), N.product(nx, ny))
}
}
val lhs: MN[F[B]] = fa.traverse[MN, B](a => (f(a), g(a)))
val rhs: MN[F[B]] = (fa.traverse(f), fa.traverse(g))
Expand Down

0 comments on commit 05f074f

Please sign in to comment.