# HG changeset patch # User haftmann # Date 1207721411 -7200 # Node ID d83271bfaba5437fcb98d29c92b6507682f15b4b # Parent 58fb6e033c00d9a276d322656c9adce0c8593317 removed syntax from monad combinators; renamed mbind to scomp diff -r 58fb6e033c00 -r d83271bfaba5 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Wed Apr 09 08:10:09 2008 +0200 +++ b/src/HOL/Fun.thy Wed Apr 09 08:10:11 2008 +0200 @@ -82,19 +82,13 @@ by (unfold comp_def, blast) -subsection {* The Forward Composition Operator @{text "f \> g"} *} +subsection {* The Forward Composition Operator @{text fcomp} *} definition fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "o>" 60) where "f o> g = (\x. g (f x))" -notation (xsymbols) - fcomp (infixl "\>" 60) - -notation (HTML output) - fcomp (infixl "\>" 60) - lemma fcomp_apply: "(f o> g) x = g (f x)" by (simp add: fcomp_def) @@ -107,6 +101,8 @@ lemma fcomp_id [simp]: "f o> id = f" by (simp add: fcomp_def) +no_notation fcomp (infixl "o>" 60) + subsection {* Injectivity and Surjectivity *} diff -r 58fb6e033c00 -r d83271bfaba5 src/HOL/Library/State_Monad.thy --- a/src/HOL/Library/State_Monad.thy Wed Apr 09 08:10:09 2008 +0200 +++ b/src/HOL/Library/State_Monad.thy Wed Apr 09 08:10:11 2008 +0200 @@ -56,29 +56,18 @@ we use a set of monad combinators: *} -definition - mbind :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" - (infixl ">>=" 60) where - "f >>= g = split g \ f" +notation fcomp (infixl ">>" 60) +notation (xsymbols) fcomp (infixl "\" 60) +notation scomp (infixl ">>=" 60) +notation (xsymbols) scomp (infixl "\=" 60) -definition - fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" - (infixl ">>" 60) where - "f >> g = g \ f" +abbreviation (input) + "return \ Pair" definition run :: "('a \ 'b) \ 'a \ 'b" where "run f = f" -syntax (xsymbols) - mbind :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" - (infixl "\=" 60) - fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" - (infixl "\" 60) - -abbreviation (input) - "return \ Pair" - print_ast_translation {* [(@{const_syntax "run"}, fn (f::ts) => Syntax.mk_appl f ts)] *} @@ -142,12 +131,12 @@ *} lemma - return_mbind [simp]: "return x \= f = f x" - unfolding mbind_def by (simp add: expand_fun_eq) + return_scomp [simp]: "return x \= f = f x" + unfolding scomp_def by (simp add: expand_fun_eq) lemma - mbind_return [simp]: "x \= return = x" - unfolding mbind_def by (simp add: expand_fun_eq split_Pair) + scomp_return [simp]: "x \= return = x" + unfolding scomp_def by (simp add: expand_fun_eq split_Pair) lemma id_fcomp [simp]: "id \ f = f" @@ -158,30 +147,30 @@ unfolding fcomp_def by simp lemma - mbind_mbind [simp]: "(f \= g) \= h = f \= (\x. g x \= h)" - unfolding mbind_def by (simp add: split_def expand_fun_eq) + scomp_scomp [simp]: "(f \= g) \= h = f \= (\x. g x \= h)" + unfolding scomp_def by (simp add: split_def expand_fun_eq) lemma - mbind_fcomp [simp]: "(f \= g) \ h = f \= (\x. g x \ h)" - unfolding mbind_def fcomp_def by (simp add: split_def expand_fun_eq) + 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_mbind [simp]: "(f \ g) \= h = f \ (g \= h)" - unfolding mbind_def fcomp_def by (simp add: split_def expand_fun_eq) + 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_mbind mbind_return id_fcomp fcomp_id - mbind_mbind mbind_fcomp fcomp_mbind fcomp_fcomp +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 - mbind_def fcomp_def run_def + scomp_def fcomp_def run_def subsection {* ML abstract operations *} @@ -194,7 +183,7 @@ fun return T sT x = Const (@{const_name return}, T --> liftT T sT) $ x; -fun mbind T1 T2 sT f g = Const (@{const_name mbind}, +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; @@ -216,7 +205,7 @@ syntax "_do" :: "do_expr \ 'a" ("do _ done" [12] 12) - "_mbind" :: "pttrn \ 'a \ do_expr \ do_expr" + "_scomp" :: "pttrn \ 'a \ do_expr \ do_expr" ("_ <- _;// _" [1000, 13, 12] 12) "_fcomp" :: "'a \ do_expr \ do_expr" ("_;// _" [13, 12] 12) @@ -226,12 +215,12 @@ ("_" [12] 12) syntax (xsymbols) - "_mbind" :: "pttrn \ 'a \ do_expr \ do_expr" + "_scomp" :: "pttrn \ 'a \ do_expr \ do_expr" ("_ \ _;// _" [1000, 13, 12] 12) translations "_do f" => "CONST run f" - "_mbind x f g" => "f \= (\x. g)" + "_scomp x f g" => "f \= (\x. g)" "_fcomp f g" => "f \ g" "_let x t f" => "CONST Let t (\x. f)" "_nil f" => "f" @@ -246,10 +235,10 @@ let val (v, t) = Syntax.variant_abs ("", dummyT, t $ Bound 0); in ((v, dummyT), t) end - fun unfold_monad (Const (@{const_syntax mbind}, _) $ f $ g) = + fun unfold_monad (Const (@{const_syntax scomp}, _) $ f $ g) = let val ((v, ty), g') = dest_abs_eta g; - in Const ("_mbind", dummyT) $ Free (v, ty) $ f $ unfold_monad g' end + 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) = diff -r 58fb6e033c00 -r d83271bfaba5 src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Wed Apr 09 08:10:09 2008 +0200 +++ b/src/HOL/Product_Type.thy Wed Apr 09 08:10:11 2008 +0200 @@ -699,34 +699,33 @@ The composition-uncurry combinator. *} -definition - mbind :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" (infixl "o->" 60) -where - "f o-> g = (\x. split g (f x))" +notation fcomp (infixl "o>" 60) -notation (xsymbols) - mbind (infixl "\\" 60) +definition + scomp :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" (infixl "o\" 60) +where + "f o\ g = (\x. split g (f x))" -notation (HTML output) - mbind (infixl "\\" 60) +lemma scomp_apply: "(f o\ g) x = split g (f x)" + by (simp add: scomp_def) -lemma mbind_apply: "(f \\ g) x = split g (f x)" - by (simp add: mbind_def) +lemma Pair_scomp: "Pair x o\ f = f x" + by (simp add: expand_fun_eq scomp_apply) -lemma Pair_mbind: "Pair x \\ f = f x" - by (simp add: expand_fun_eq mbind_apply) +lemma scomp_Pair: "x o\ Pair = x" + by (simp add: expand_fun_eq scomp_apply) -lemma mbind_Pair: "x \\ Pair = x" - by (simp add: expand_fun_eq mbind_apply) +lemma scomp_scomp: "(f o\ g) o\ h = f o\ (\x. g x o\ h)" + by (simp add: expand_fun_eq split_twice scomp_def) -lemma mbind_mbind: "(f \\ g) \\ h = f \\ (\x. g x \\ h)" - by (simp add: expand_fun_eq split_twice mbind_def) +lemma scomp_fcomp: "(f o\ g) o> h = f o\ (\x. g x o> h)" + by (simp add: expand_fun_eq scomp_apply fcomp_def split_def) -lemma mbind_fcomp: "(f \\ g) \> h = f \\ (\x. g x \> h)" - by (simp add: expand_fun_eq mbind_apply fcomp_def split_def) +lemma fcomp_scomp: "(f o> g) o\ h = f o> (g o\ h)" + by (simp add: expand_fun_eq scomp_apply fcomp_apply) -lemma fcomp_mbind: "(f \> g) \\ h = f \> (g \\ h)" - by (simp add: expand_fun_eq mbind_apply fcomp_apply) +no_notation fcomp (infixl "o>" 60) +no_notation scomp (infixl "o\" 60) text {*