(* Title: HOL/Library/State_Monad.thy ID: $Id: State_Monad.thy,v 1.19 2008/04/09 06:10:11 haftmann Exp $ Author: Florian Haftmann, TU Muenchen *) header {* Combinators syntax for generic, open state monads (single threaded monads) *} theory State_Monad imports List begin subsection {* Motivation *} text {* The logic HOL has no notion of constructor classes, so it is not possible to model monads the Haskell way in full genericity in Isabelle/HOL. However, this theory provides substantial support for a very common class of monads: \emph{state monads} (or \emph{single-threaded monads}, since a state is transformed single-threaded). To enter from the Haskell world, \url{http://www.engr.mun.ca/~theo/Misc/haskell_and_monads.htm} makes a good motivating start. Here we just sketch briefly how those monads enter the game of Isabelle/HOL. *} subsection {* State transformations and combinators *} text {* We classify functions operating on states into two categories: \begin{description} \item[transformations] with type signature @{text "σ => σ'"}, transforming a state. \item[``yielding'' transformations] with type signature @{text "σ => α × σ'"}, ``yielding'' a side result while transforming a state. \item[queries] with type signature @{text "σ => α"}, computing a result dependent on a state. \end{description} By convention we write @{text "σ"} for types representing states and @{text "α"}, @{text "β"}, @{text "γ"}, @{text "…"} for types representing side results. Type changes due to transformations are not excluded in our scenario. We aim to assert that values of any state type @{text "σ"} are used in a single-threaded way: after application of a transformation on a value of type @{text "σ"}, the former value should not be used again. To achieve this, we use a set of monad combinators: *} notation fcomp (infixl ">>" 60) notation (xsymbols) fcomp (infixl "»" 60) notation scomp (infixl ">>=" 60) notation (xsymbols) scomp (infixl "»=" 60) abbreviation (input) "return ≡ Pair" definition run :: "('a => 'b) => 'a => 'b" where "run f = f" print_ast_translation {* [(@{const_syntax "run"}, fn (f::ts) => Syntax.mk_appl f ts)] *} text {* Given two transformations @{term f} and @{term g}, they may be directly composed using the @{term "op »"} combinator, forming a forward composition: @{prop "(f » g) s = f (g s)"}. After any yielding transformation, we bind the side result immediately using a lambda abstraction. This is the purpose of the @{term "op »="} combinator: @{prop "(f »= (λx. g)) s = (let (x, s') = f s in g s')"}. For queries, the existing @{term "Let"} is appropriate. Naturally, a computation may yield a side result by pairing it to the state from the left; we introduce the suggestive abbreviation @{term return} for this purpose. The @{const run} ist just a marker. The most crucial distinction to Haskell is that we do not need to introduce distinguished type constructors for different kinds of state. This has two consequences: \begin{itemize} \item The monad model does not state anything about the kind of state; the model for the state is completely orthogonal and may be specified completely independently. \item There is no distinguished type constructor encapsulating away the state transformation, i.e.~transformations may be applied directly without using any lifting or providing and dropping units (``open monad''). \item The type of states may change due to a transformation. \end{itemize} *} subsection {* Obsolete runs *} text {* @{term run} is just a doodle and should not occur nested: *} lemma run_simp [simp]: "!!f. run (run f) = run f" "!!f g. run f »= g = f »= g" "!!f g. run f » g = f » g" "!!f g. f »= (λx. run g) = f »= (λx. g)" "!!f g. f » run g = f » g" "!!f. f = run f <-> True" "!!f. run f = f <-> True" unfolding run_def by rule+ subsection {* Monad laws *} text {* The common monadic laws hold and may also be used as normalization rules for monadic expressions: *} lemma return_scomp [simp]: "return x »= f = f x" unfolding scomp_def by (simp add: expand_fun_eq) lemma scomp_return [simp]: "x »= return = x" unfolding scomp_def by (simp add: expand_fun_eq split_Pair) lemma id_fcomp [simp]: "id » f = f" unfolding fcomp_def by simp lemma fcomp_id [simp]: "f » id = f" unfolding fcomp_def by simp lemma scomp_scomp [simp]: "(f »= g) »= h = f »= (λx. g x »= h)" unfolding scomp_def by (simp add: split_def expand_fun_eq) lemma scomp_fcomp [simp]: "(f »= g) » h = f »= (λx. g x » h)" unfolding scomp_def fcomp_def by (simp add: split_def expand_fun_eq) lemma fcomp_scomp [simp]: "(f » g) »= h = f » (g »= h)" unfolding scomp_def fcomp_def by (simp add: split_def expand_fun_eq) lemma fcomp_fcomp [simp]: "(f » g) » h = f » (g » h)" unfolding fcomp_def o_assoc .. lemmas monad_simp = run_simp return_scomp scomp_return id_fcomp fcomp_id scomp_scomp scomp_fcomp fcomp_scomp fcomp_fcomp text {* Evaluation of monadic expressions by force: *} lemmas monad_collapse = monad_simp o_apply o_assoc split_Pair split_comp scomp_def fcomp_def run_def subsection {* ML abstract operations *} ML {* structure StateMonad = struct fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); fun liftT' sT = sT --> sT; fun return T sT x = Const (@{const_name return}, T --> liftT T sT) $ x; fun scomp T1 T2 sT f g = Const (@{const_name scomp}, liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; fun run T sT f = Const (@{const_name run}, liftT' (liftT T sT)) $ f; end; *} subsection {* Syntax *} text {* We provide a convenient do-notation for monadic expressions well-known from Haskell. @{const Let} is printed specially in do-expressions. *} nonterminals do_expr syntax "_do" :: "do_expr => 'a" ("do _ done" [12] 12) "_scomp" :: "pttrn => 'a => do_expr => do_expr" ("_ <- _;// _" [1000, 13, 12] 12) "_fcomp" :: "'a => do_expr => do_expr" ("_;// _" [13, 12] 12) "_let" :: "pttrn => 'a => do_expr => do_expr" ("let _ = _;// _" [1000, 13, 12] 12) "_nil" :: "'a => do_expr" ("_" [12] 12) syntax (xsymbols) "_scomp" :: "pttrn => 'a => do_expr => do_expr" ("_ \<leftarrow> _;// _" [1000, 13, 12] 12) translations "_do f" => "CONST run f" "_scomp x f g" => "f »= (λx. g)" "_fcomp f g" => "f » g" "_let x t f" => "CONST Let t (λx. f)" "_nil f" => "f" print_translation {* let fun dest_abs_eta (Abs (abs as (_, ty, _))) = let val (v, t) = Syntax.variant_abs abs; in ((v, ty), t) end | dest_abs_eta t = let val (v, t) = Syntax.variant_abs ("", dummyT, t $ Bound 0); in ((v, dummyT), t) end fun unfold_monad (Const (@{const_syntax scomp}, _) $ f $ g) = let val ((v, ty), g') = dest_abs_eta g; in Const ("_scomp", dummyT) $ Free (v, ty) $ f $ unfold_monad g' end | unfold_monad (Const (@{const_syntax fcomp}, _) $ f $ g) = Const ("_fcomp", dummyT) $ f $ unfold_monad g | unfold_monad (Const (@{const_syntax Let}, _) $ f $ g) = let val ((v, ty), g') = dest_abs_eta g; in Const ("_let", dummyT) $ Free (v, ty) $ f $ unfold_monad g' end | unfold_monad (Const (@{const_syntax Pair}, _) $ f) = Const ("return", dummyT) $ f | unfold_monad f = f; fun tr' (f::ts) = list_comb (Const ("_do", dummyT) $ unfold_monad f, ts) in [(@{const_syntax "run"}, tr')] end; *} subsection {* Combinators *} definition lift :: "('a => 'b) => 'a => 'c => 'b × 'c" where "lift f x = return (f x)" primrec list :: "('a => 'b => 'b) => 'a list => 'b => 'b" where "list f [] = id" | "list f (x#xs) = (do f x; list f xs done)" primrec list_yield :: "('a => 'b => 'c × 'b) => 'a list => 'b => 'c list × 'b" where "list_yield f [] = return []" | "list_yield f (x#xs) = (do y \<leftarrow> f x; ys \<leftarrow> list_yield f xs; return (y#ys) done)" definition collapse :: "('a => ('a => 'b × 'a) × 'a) => 'a => 'b × 'a" where "collapse f = (do g \<leftarrow> f; g done)" text {* combinator properties *} lemma list_append [simp]: "list f (xs @ ys) = list f xs » list f ys" by (induct xs) (simp_all del: id_apply) lemma list_cong [fundef_cong, recdef_cong]: "[| !!x. x ∈ set xs ==> f x = g x; xs = ys |] ==> list f xs = list g ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) from Cons have "!!y. y ∈ set (x # xs) ==> f y = g y" by auto then have "!!y. y ∈ set xs ==> f y = g y" by auto with Cons have "list f xs = list g xs" by auto with Cons have "list f (x # xs) = list g (x # xs)" by auto with Cons show "list f (x # xs) = list g ys" by auto qed lemma list_yield_cong [fundef_cong, recdef_cong]: "[| !!x. x ∈ set xs ==> f x = g x; xs = ys |] ==> list_yield f xs = list_yield g ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) from Cons have "!!y. y ∈ set (x # xs) ==> f y = g y" by auto then have "!!y. y ∈ set xs ==> f y = g y" by auto with Cons have "list_yield f xs = list_yield g xs" by auto with Cons have "list_yield f (x # xs) = list_yield g (x # xs)" by auto with Cons show "list_yield f (x # xs) = list_yield g ys" by auto qed text {* For an example, see HOL/ex/Random.thy. *} end
lemma run_simp:
f = f
f >>= g = f >>= g
f >> g = f >> g
f >>= (λx. g) = f >>= (λx. g)
f >> g = f >> g
(f = f) = True
(f = f) = True
lemma return_scomp:
Pair x >>= f = f x
lemma scomp_return:
x >>= Pair = x
lemma id_fcomp:
id >> f = f
lemma fcomp_id:
f >> id = f
lemma scomp_scomp:
f >>= g >>= h = f >>= (λx. g x >>= h)
lemma scomp_fcomp:
f >>= g >> h = f >>= (λx. g x >> h)
lemma fcomp_scomp:
f >> g >>= h = f >> (g >>= h)
lemma fcomp_fcomp:
f >> g >> h = f >> (g >> h)
lemma monad_simp:
f = f
f >>= g = f >>= g
f >> g = f >> g
f >>= (λx. g) = f >>= (λx. g)
f >> g = f >> g
(f = f) = True
(f = f) = True
Pair x >>= f = f x
x >>= Pair = x
id >> f = f
f >> id = f
f >>= g >>= h = f >>= (λx. g x >>= h)
f >>= g >> h = f >>= (λx. g x >> h)
f >> g >>= h = f >> (g >>= h)
f >> g >> h = f >> (g >> h)
lemma monad_collapse:
f = f
f >>= g = f >>= g
f >> g = f >> g
f >>= (λx. g) = f >>= (λx. g)
f >> g = f >> g
(f = f) = True
(f = f) = True
Pair x >>= f = f x
x >>= Pair = x
id >> f = f
f >> id = f
f >>= g >>= h = f >>= (λx. g x >>= h)
f >>= g >> h = f >>= (λx. g x >> h)
f >> g >>= h = f >> (g >>= h)
f >> g >> h = f >> (g >> h)
(f o g) x = f (g x)
f o (g o h) = f o g o h
(λ(x, y). (x, y)) = id
split (f o g) x = f (g (fst x)) (snd x)
f >>= g = (λx. split g (f x))
f >> g = (λx. g (f x))
f = f
lemma list_append:
list f (xs @ ys) = list f xs >> list f ys
lemma list_cong:
[| !!x. x ∈ set xs ==> f x = g x; xs = ys |] ==> list f xs = list g ys
lemma list_yield_cong:
[| !!x. x ∈ set xs ==> f x = g x; xs = ys |]
==> list_yield f xs = list_yield g ys