Documentation

Std.Do.WP.SimpLemmas

Simp lemmas for working with weakest preconditions #

Many weakest preconditions are so simple that we can compute them directly or express them in terms of "simpler" weakest preconditions. This module provides simp lemmas targeting expressions of the form wp⟦x⟧ Q and rewrites them to expressions of simpler weakest preconditions.

WP #

@[simp]
theorem Std.Do.WP.ReaderT_run {m : TypeType u} {ps : PostShape} {ρ α : Type} {r : ρ} {Q : PostCond α ps} [WP m ps] (x : ReaderT ρ m α) :
wp⟦x.run r Q = wp⟦x (fun (a : α) (x : ρ) => Q.fst a, Q.snd) r
@[simp]
theorem Std.Do.WP.StateT_run {m : TypeType u} {ps : PostShape} {σ α : Type} {s : σ} {Q : PostCond (α × σ) ps} [WP m ps] (x : StateT σ m α) :
wp⟦x.run s Q = wp⟦x (fun (a : α) (s : σ) => Q.fst (a, s), Q.snd) s
@[simp]
theorem Std.Do.WP.ExceptT_run {m : TypeType u} {ps : PostShape} {ε α : Type} {Q : PostCond (Except ε α) ps} [WP m ps] (x : ExceptT ε m α) :
wp⟦x.run Q = wp⟦x (fun (a : α) => Q.fst (Except.ok a), fun (e : ε) => Q.fst (Except.error e), Q.snd)

WPMonad #

@[simp]
theorem Std.Do.WP.pure {m : TypeType u} {ps : PostShape} {α : Type} [Monad m] [WPMonad m ps] (a : α) (Q : PostCond α ps) :
@[simp]
theorem Std.Do.WP.bind {m : TypeType u} {ps : PostShape} {α β : Type} [Monad m] [WPMonad m ps] (x : m α) (f : αm β) (Q : PostCond β ps) :
wp⟦x >>= f Q = wp⟦x (fun (a : α) => wp⟦f a Q, Q.snd)
@[simp]
theorem Std.Do.WP.map {m : TypeType u} {ps : PostShape} {α β : Type} [Monad m] [WPMonad m ps] (f : αβ) (x : m α) (Q : PostCond β ps) :
wp⟦f <$> x Q = wp⟦x (fun (a : α) => Q.fst (f a), Q.snd)
@[simp]
theorem Std.Do.WP.seq {m : TypeType u} {ps : PostShape} {α β : Type} [Monad m] [WPMonad m ps] (f : m (αβ)) (x : m α) (Q : PostCond β ps) :
wp⟦f <*> x Q = wp⟦f (fun (f : αβ) => wp⟦x (fun (a : α) => Q.fst (f a), Q.snd), Q.snd)

MonadLift #

The definitions that follow interpret liftM and thus instances of, e.g., MonadStateOf.

@[simp]
theorem Std.Do.WP.monadLift_StateT {m : TypeType u} {ps : PostShape} {α σ : Type} [Monad m] [WPMonad m ps] (x : m α) (Q : PostCond α (PostShape.arg σ ps)) :
wp⟦MonadLift.monadLift x Q = fun (s : σ) => wp⟦x (fun (a : α) => Q.fst a s, Q.snd)
@[simp]
theorem Std.Do.WP.monadLift_ReaderT {m : TypeType u} {ps : PostShape} {α ρ : Type} [Monad m] [WPMonad m ps] (x : m α) (Q : PostCond α (PostShape.arg ρ ps)) :
wp⟦MonadLift.monadLift x Q = fun (s : ρ) => wp⟦x (fun (a : α) => Q.fst a s, Q.snd)
@[simp]
theorem Std.Do.WP.monadLift_ExceptT {m : TypeType u} {ps : PostShape} {α ε : Type} [Monad m] [WPMonad m ps] (x : m α) (Q : PostCond α (PostShape.except ε ps)) :
wp⟦MonadLift.monadLift x Q = wp⟦x (fun (a : α) => Q.fst a, Q.snd.snd)
@[simp]
theorem Std.Do.WP.monadLift_trans {m : TypeType u} {ps : PostShape} {o : TypeType u_1} {n : TypeType u_2} {α : Type} {x : m α} {Q : PostCond α ps} [WP o ps] [MonadLift n o] [MonadLiftT m n] :
@[simp]
theorem Std.Do.WP.monadLift_refl {m : TypeType u} {ps : PostShape} {α : Type} {x : m α} {Q : PostCond α ps} [WP m ps] :
@[simp]
theorem Std.Do.WP.lift_StateT {m : TypeType u} {ps : PostShape} {α σ : Type} {Q : PostCond α (PostShape.arg σ ps)} [WP m ps] [Monad m] (x : m α) :
@[simp]
theorem Std.Do.WP.lift_ExceptT {m : TypeType u} {ps : PostShape} {α ε : Type} {Q : PostCond α (PostShape.except ε ps)} [WP m ps] [Monad m] (x : m α) :
@[simp]
theorem Std.Do.WP.read_MonadReaderOf {m : TypeType u} {ρ : Type} {n : TypeType u_1} {s✝ : PostShape} {Q : PostCond ρ s✝} [MonadReaderOf ρ m] [MonadLift m n] [WP n s✝] :
@[simp]
theorem Std.Do.WP.readThe {m : TypeType u} {ps : PostShape} {ρ : Type} {Q : PostCond ρ ps} [MonadReaderOf ρ m] [WP m ps] :
@[simp]
theorem Std.Do.WP.read_MonadReader {m : TypeType u} {ps : PostShape} {ρ : Type} {Q : PostCond ρ ps} [MonadReaderOf ρ m] [WP m ps] :
@[simp]
theorem Std.Do.WP.get_MonadStateOf {m : TypeType u} {σ : Type} {n : TypeType u_1} {s✝ : PostShape} {Q : PostCond σ s✝} [MonadStateOf σ m] [MonadLift m n] [WP n s✝] :
@[simp]
theorem Std.Do.WP.get_MonadState {m : TypeType u} {ps : PostShape} {σ : Type} {Q : PostCond σ ps} [WP m ps] [MonadStateOf σ m] :
@[simp]
theorem Std.Do.WP.getThe_MonadStateOf {m : TypeType u} {ps : PostShape} {σ : Type} {Q : PostCond σ ps} [WP m ps] [MonadStateOf σ m] :
@[simp]
theorem Std.Do.WP.set_MonadStateOf {m : TypeType u} {σ : Type} {n : TypeType u_1} {x : σ} {s✝ : PostShape} {Q : PostCond PUnit s✝} [MonadStateOf σ m] [MonadLift m n] [WP n s✝] :
@[simp]
theorem Std.Do.WP.set_MonadState {m : TypeType u} {ps : PostShape} {σ : Type} {x : σ} {Q : PostCond PUnit ps} [WP m ps] [MonadStateOf σ m] :
@[simp]
theorem Std.Do.WP.modifyGet_MonadStateOf {m : TypeType u} {σ : Type} {n : TypeType u_1} {α : Type} {s✝ : PostShape} {Q : PostCond α s✝} [MonadStateOf σ m] [MonadLift m n] [WP n s✝] (f : σα × σ) :
@[simp]
theorem Std.Do.WP.modifyGet_MonadState {m : TypeType u} {ps : PostShape} {σ α : Type} {Q : PostCond α ps} [WP m ps] [MonadStateOf σ m] (f : σα × σ) :
@[simp]
theorem Std.Do.WP.modifyGetThe_MonadStateOf {m : TypeType u} {ps : PostShape} {σ α : Type} {Q : PostCond α ps} [WP m ps] [MonadStateOf σ m] (f : σα × σ) :
@[simp]
theorem Std.Do.WP.modify_MonadStateOf {m : TypeType u} {ps : PostShape} {σ : Type} {Q : PostCond PUnit ps} [WP m ps] [MonadStateOf σ m] (f : σσ) :
@[simp]
theorem Std.Do.WP.modifyThe_MonadStateOf {m : TypeType u} {ps : PostShape} {σ : Type} {Q : PostCond PUnit ps} [WP m ps] [MonadStateOf σ m] (f : σσ) :
@[simp]
theorem Std.Do.WP.read_ReaderT {m : TypeType u} {ps : PostShape} {ρ : Type} {Q : PostCond ρ (PostShape.arg ρ ps)} [Monad m] [WPMonad m ps] :
wp⟦MonadReaderOf.read Q = fun (r : ρ) => Q.fst r r
@[simp]
theorem Std.Do.WP.get_StateT {m : TypeType u} {ps : PostShape} {σ : Type} {Q : PostCond σ (PostShape.arg σ ps)} [Monad m] [WPMonad m ps] :
wp⟦MonadStateOf.get Q = fun (s : σ) => Q.fst s s
@[simp]
theorem Std.Do.WP.set_StateT {m : TypeType u} {ps : PostShape} {σ : Type} {Q : PostCond PUnit (PostShape.arg σ ps)} [Monad m] [WPMonad m ps] (x : σ) :
wp⟦set x Q = fun (x_1 : σ) => Q.fst PUnit.unit x
@[simp]
theorem Std.Do.WP.modifyGet_StateT {m : TypeType u} {ps : PostShape} {σ α : Type} {Q : PostCond α (PostShape.arg σ ps)} [Monad m] [WPMonad m ps] (f : σα × σ) :
wp⟦MonadStateOf.modifyGet f Q = fun (s : σ) => Q.fst (f s).fst (f s).snd
@[simp]
theorem Std.Do.WP.get_EStateM {ε σ : Type} {Q : PostCond σ (PostShape.except ε (PostShape.arg σ PostShape.pure))} :
wp⟦MonadStateOf.get Q = fun (s : σ) => Q.fst s s
@[simp]
theorem Std.Do.WP.set_EStateM {σ ε : Type} {Q : PostCond PUnit (PostShape.except ε (PostShape.arg σ PostShape.pure))} (x : σ) :
wp⟦set x Q = fun (x_1 : σ) => Q.fst PUnit.unit x
@[simp]
theorem Std.Do.WP.modifyGet_EStateM {σ α ε : Type} {Q : PostCond α (PostShape.except ε (PostShape.arg σ PostShape.pure))} (f : σα × σ) :
wp⟦MonadStateOf.modifyGet f Q = fun (s : σ) => Q.fst (f s).fst (f s).snd

MonadFunctor #

The definitions that follow interpret monadMap and thus instances of, e.g., MonadReaderWithOf.

@[simp]
theorem Std.Do.WP.monadMap_StateT {ps : PostShape} {σ : Type} (m : TypeType u) [Monad m] [WP m ps] (f : {β : Type} → m βm β) {α : Type} (x : StateT σ m α) (Q : PostCond α (PostShape.arg σ ps)) :
wp⟦MonadFunctor.monadMap (fun {β : Type} => f) x Q = fun (s : σ) => wp⟦f (x.run s) (fun (x : α × σ) => match x with | (a, s) => Q.fst a s, Q.snd)
@[simp]
theorem Std.Do.WP.monadMap_ReaderT {ps : PostShape} {ρ : Type} (m : TypeType u) [Monad m] [WP m ps] (f : {β : Type} → m βm β) {α : Type} (x : ReaderT ρ m α) (Q : PostCond α (PostShape.arg ρ ps)) :
wp⟦MonadFunctor.monadMap (fun {β : Type} => f) x Q = fun (s : ρ) => wp⟦f (x.run s) (fun (a : α) => Q.fst a s, Q.snd)
@[simp]
theorem Std.Do.WP.monadMap_ExceptT {ps : PostShape} {ε : Type} (m : TypeType u) [Monad m] [WP m ps] (f : {β : Type} → m βm β) {α : Type} (x : ExceptT ε m α) (Q : PostCond α (PostShape.except ε ps)) :
wp⟦MonadFunctor.monadMap (fun {β : Type} => f) x Q = wp⟦f x.run (fun (x : Except ε α) => match x with | Except.ok a => Q.fst a | Except.error e => Q.snd.fst e, Q.snd.snd)
@[simp]
theorem Std.Do.WP.monadMap_trans {m : TypeType u} {ps : PostShape} {o : TypeType u_1} {n : TypeType u_2} {α : Type} {f : {β : Type} → m βm β} {x : o α} {Q : PostCond α ps} [WP o ps] [MonadFunctor n o] [MonadFunctorT m n] :
wp⟦monadMap f x Q = wp⟦MonadFunctor.monadMap (fun {β : Type} => monadMap fun {β : Type} => f) x Q
@[simp]
theorem Std.Do.WP.monadMap_refl {m : TypeType u} {ps : PostShape} {α : Type} {f : {β : Type} → m βm β} {x : m α} {Q : PostCond α ps} [WP m ps] :
@[simp]
theorem Std.Do.WP.withReader_ReaderT {m : TypeType u} {ps : PostShape} {ρ α : Type} {f : ρρ} {x : ReaderT ρ m α} {Q : PostCond α (PostShape.arg ρ ps)} [WP m ps] :
wp⟦MonadWithReaderOf.withReader f x Q = fun (r : ρ) => wp⟦x (fun (a : α) (x : ρ) => Q.fst a r, Q.snd) (f r)
@[simp]
theorem Std.Do.WP.withReader_MonadWithReaderOf {m : TypeType u} {ρ : Type} {n : TypeType u} {nsh : PostShape} {α : Type} {Q : PostCond α nsh} [MonadWithReaderOf ρ m] [WP n nsh] [MonadFunctor m n] (f : ρρ) (x : n α) :
@[simp]
theorem Std.Do.WP.withReader_MonadWithReader {m : TypeType u} {ps : PostShape} {ρ α : Type} {Q : PostCond α ps} [MonadWithReaderOf ρ m] [WP m ps] (f : ρρ) (x : m α) :
@[simp]
theorem Std.Do.WP.withTheReader {m : TypeType u} {ps : PostShape} {ρ α : Type} {Q : PostCond α ps} [MonadWithReaderOf ρ m] [WP m ps] (f : ρρ) (x : m α) :

MonadExceptOf #

The definitions that follow interpret throw, throwThe, tryCatch, etc.

Since MonadExceptOf cannot be lifted through MonadLift or MonadFunctor, there is one lemma per MonadExceptOf operation and instance to lift through; the classic m*n instances problem.

@[simp]
theorem Std.Do.WP.throw_MonadExcept {m : TypeType u} {ps : PostShape} {ε : Type u_1} {α : Type} {e : ε} {Q : PostCond α ps} [MonadExceptOf ε m] [WP m ps] :
@[simp]
theorem Std.Do.WP.throwThe {m : TypeType u} {ps : PostShape} {ε : Type u_1} {α : Type} {e : ε} {Q : PostCond α ps} [MonadExceptOf ε m] [WP m ps] :
@[simp]
theorem Std.Do.WP.throw_ExceptT {m : TypeType u} {ps : PostShape} {ε α : Type} {e : ε} {Q : PostCond α (PostShape.except ε ps)} [Monad m] [WPMonad m ps] :
@[simp]
theorem Std.Do.WP.throw_ReaderT {m : TypeType u} {sh : PostShape} {ε : Type u_1} {ρ α : Type} {e : ε} {Q : PostCond α (PostShape.arg ρ sh)} [WP m sh] [Monad m] [MonadExceptOf ε m] :
@[simp]
theorem Std.Do.WP.throw_StateT {m : TypeType u} {sh : PostShape} {ε : Type u_1} {ρ α : Type} {e : ε} {Q : PostCond α (PostShape.arg ρ sh)} [WP m sh] [Monad m] [MonadExceptOf ε m] :
@[simp]
theorem Std.Do.WP.throw_lift_ExceptT {m : TypeType u} {sh : PostShape} {ε ε' α : Type} {e : ε} {Q : PostCond α (PostShape.except ε' sh)} [WP m sh] [Monad m] [MonadExceptOf ε m] :
wp⟦MonadExceptOf.throw e Q = wp⟦MonadExceptOf.throw e (fun (x : Except ε' α) => match x with | Except.ok a => Q.fst a | Except.error e => Q.snd.fst e, Q.snd.snd)
@[simp]
theorem Std.Do.WP.tryCatch_MonadExcept {m : TypeType u} {ps : PostShape} {ε : Type u_1} {α : Type} {x : m α} {h : εm α} {Q : PostCond α ps} [MonadExceptOf ε m] [WP m ps] :
@[simp]
theorem Std.Do.WP.tryCatchThe {m : TypeType u} {ps : PostShape} {ε : Type u_1} {α : Type} {x : m α} {h : εm α} {Q : PostCond α ps} [MonadExceptOf ε m] [WP m ps] :
@[simp]
theorem Std.Do.WP.tryCatch_Except {ε α : Type} {x : Except ε α} {h : εExcept ε α} {Q : PostCond α (PostShape.except ε PostShape.pure)} :
@[simp]
theorem Std.Do.WP.tryCatch_ExceptT {m : TypeType u} {ps : PostShape} {ε α : Type} {x : ExceptT ε m α} {h : εExceptT ε m α} {Q : PostCond α (PostShape.except ε ps)} [Monad m] [WPMonad m ps] :
@[simp]
theorem Std.Do.WP.tryCatch_EStateM {ε σ δ α : Type} {x : EStateM ε σ α} {h : εEStateM ε σ α} {Q : PostCond α (PostShape.except ε (PostShape.arg σ PostShape.pure))} [EStateM.Backtrackable δ σ] :
wp⟦MonadExceptOf.tryCatch x h Q = fun (s : σ) => wp⟦x (Q.fst, fun (e : ε) (s' : σ) => wp⟦h e Q (EStateM.Backtrackable.restore s' (EStateM.Backtrackable.save s)), Q.snd.snd) s
@[simp]
theorem Std.Do.WP.tryCatch_ReaderT {m : TypeType u} {sh : PostShape} {ε : Type u_1} {ρ α : Type} {x : ReaderT ρ m α} {h : εReaderT ρ m α} {Q : PostCond α (PostShape.arg ρ sh)} [WP m sh] [Monad m] [MonadExceptOf ε m] :
wp⟦MonadExceptOf.tryCatch x h Q = fun (r : ρ) => wp⟦MonadExceptOf.tryCatch (x.run r) fun (e : ε) => (h e).run r (fun (a : α) => Q.fst a r, Q.snd)
@[simp]
theorem Std.Do.WP.tryCatch_StateT {m : TypeType u} {sh : PostShape} {ε : Type u_1} {σ α : Type} {x : StateT σ m α} {h : εStateT σ m α} {Q : PostCond α (PostShape.arg σ sh)} [WP m sh] [Monad m] [MonadExceptOf ε m] :
wp⟦MonadExceptOf.tryCatch x h Q = fun (s : σ) => wp⟦MonadExceptOf.tryCatch (x.run s) fun (e : ε) => (h e).run s (fun (xs : α × σ) => Q.fst xs.fst xs.snd, Q.snd)
@[simp]
theorem Std.Do.WP.tryCatch_lift_ExceptT {m : TypeType u} {sh : PostShape} {ε ε' α : Type} {x : ExceptT ε' m α} {h : εExceptT ε' m α} {Q : PostCond α (PostShape.except ε' sh)} [WP m sh] [Monad m] [MonadExceptOf ε m] :
wp⟦MonadExceptOf.tryCatch x h Q = wp⟦MonadExceptOf.tryCatch x h (fun (x : Except ε' α) => match x with | Except.ok a => Q.fst a | Except.error e => Q.snd.fst e, Q.snd.snd)