SatisfiesM #
The SatisfiesM
predicate works over an arbitrary (lawful) monad / applicative / functor,
and enables Hoare-like reasoning over monadic expressions. For example, given a monadic
function f : α → m β
, to say that the return value of f
satisfies Q
whenever
the input satisfies P
, we write ∀ a, P a → SatisfiesM Q (f a)
.
For any monad equipped with MonadSatisfying m
one can lift SatisfiesM
to a monadic value in Subtype
,
using satisfying x h : m {a // p a}
, where x : m α
and h : SatisfiesM p x
.
This includes Option
, ReaderT
, StateT
, and ExceptT
, and the Lean monad stack.
(Although it is not entirely clear one should treat the Lean monad stack as lawful,
even though Lean accepts this.)
Notes #
SatisfiesM
is not yet a satisfactory solution for verifying the behaviour of large scale monadic
programs. Such a solution would allow ergonomic reasoning about large do
blocks,
with convenient mechanisms for introducing invariants and loop conditions as needed.
It is possible that in the future SatiesfiesM
will become part of such a solution,
presumably requiring more syntactic support (and smarter do
blocks) from Lean.
Or it may be that such a solution will look different!
This is an open research program, and for now one should not be overly ambitious using SatisfiesM
.
In particular lemmas about pure operations on data structures in Batteries
except for HashMap
should avoid SatisfiesM
for now, so that it is easy to migrate to other approaches in future.
SatisfiesM p (x : m α)
lifts propositions over a monad. It asserts that x
may as well
have the type x : m {a // p a}
, because there exists some m {a // p a}
whose image is x
.
So p
is the postcondition of the monadic value.
Equations
- SatisfiesM p x = ∃ (x' : m { a : α // p a }), Subtype.val <$> x' = x
Instances For
If p
is always true, then every x
satisfies it.
If p
is always true, then every x
satisfies it.
(This is the strongest postcondition version of of_true
.)
The SatisfiesM p x
predicate is monotonic in p
.
SatisfiesM
distributes over <$>
, general version.
SatisfiesM
distributes over <$>
, strongest postcondition version.
(Use this for reasoning forward from assumptions.)
SatisfiesM
distributes over <$>
, weakest precondition version.
(Use this for reasoning backward from the goal.)
SatisfiesM
distributes over mapConst
, general version.
SatisfiesM
distributes over pure
, general version / weakest precondition version.
SatisfiesM
distributes over <*>
, general version.
SatisfiesM
distributes over <*>
, strongest postcondition version.
SatisfiesM
distributes over <*>
, weakest precondition version 1.
(Use this when x
and the goal are known and f
is a subgoal.)
SatisfiesM
distributes over <*>
, weakest precondition version 2.
(Use this when f
and the goal are known and x
is a subgoal.)
SatisfiesM
distributes over <*
, general version.
SatisfiesM
distributes over *>
, general version.
SatisfiesM
distributes over >>=
, general version.
SatisfiesM
distributes over >>=
, weakest precondition version.
If a monad has MonadSatisfying m
, then we can lift a h : SatisfiesM (m := m) p x
predicate
to monadic value satisfying x p : m { x // p x }
.
Reader, state, and exception monads have MonadSatisfying
instances if the base monad does.
- satisfying : {α : Type u} → {p : α → Prop} → {x : m α} → SatisfiesM p x → m { a : α // p a }
Lift a
SatisfiesM
predicate to a monadic value. - val_eq : ∀ {α : Type u} {p : α → Prop} {x : m α} (h : SatisfiesM p x), Subtype.val <$> satisfying h = x
The value of the lifted monadic value is equal to the original monadic value.
Instances
Equations
- MonadSatisfying.instId = { satisfying := fun {α : Type ?u.10} {p : α → Prop} {x : Id α} (h : SatisfiesM p x) => ⟨x, ⋯⟩, val_eq := @MonadSatisfying.instId.proof_3 }
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- MonadSatisfying.instStateRefT' = inferInstanceAs (MonadSatisfying (ReaderT (ST.Ref ω σ) m))
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.