# HG changeset patch # User haftmann # Date 1163805627 -3600 # Node ID 4bc2882f80af42b80afab03f83c5896e24c36137 # Parent 13c33ad15303609afec1bd09889069e066041068 added combinators and lemmas diff -r 13c33ad15303 -r 4bc2882f80af src/HOL/Library/State_Monad.thy --- a/src/HOL/Library/State_Monad.thy Sat Nov 18 00:20:26 2006 +0100 +++ b/src/HOL/Library/State_Monad.thy Sat Nov 18 00:20:27 2006 +0100 @@ -144,7 +144,6 @@ "\f. run f = f \ True" unfolding run_def by rule+ - subsection {* Monad laws *} text {* @@ -161,6 +160,14 @@ unfolding mbind_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 mbind_mbind [simp]: "(f \= g) \= h = f \= (\x. g x \= h)" unfolding mbind_def by (simp add: split_def expand_fun_eq) @@ -176,7 +183,7 @@ fcomp_fcomp [simp]: "(f \ g) \ h = f \ (g \ h)" unfolding fcomp_def o_assoc .. -lemmas monad_simp = run_simp return_mbind mbind_return +lemmas monad_simp = run_simp return_mbind mbind_return id_fcomp fcomp_id mbind_mbind mbind_fcomp fcomp_mbind fcomp_fcomp text {* @@ -246,6 +253,61 @@ ] end; *} +subsection {* Combinators *} + +definition + lift :: "('a \ 'b) \ 'a \ 'c \ 'b \ 'c" + "lift f x = return (f x)" + +fun + list :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where + "list f [] = id" + "list f (x#xs) = (do f x; list f xs done)" +lemmas [code] = list.simps + +fun list_yield :: "('a \ 'b \ 'c \ 'b) \ 'a list \ 'b \ 'c list \ 'b" where + "list_yield f [] = return []" + "list_yield f (x#xs) = (do y \ f x; ys \ list_yield f xs; return (y#ys) done)" +lemmas [code] = list_yield.simps + +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) (*FIXME*) + +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 f xs arbitrary: g ys rule: list.induct) + case 1 then show ?case by simp +next + case (2 f x xs g) + from 2 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 2 have "list f xs = list g xs" by auto + with 2 have "list f (x # xs) = list g (x # xs)" by auto + with 2 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 f xs arbitrary: g ys rule: list_yield.induct) + case 1 then show ?case by simp +next + case (2 f x xs g) + from 2 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 2 have "list_yield f xs = list_yield g xs" by auto + with 2 have "list_yield f (x # xs) = list_yield g (x # xs)" by auto + with 2 show "list_yield f (x # xs) = list_yield g ys" by auto +qed + +text {* + still waiting for extensions@{text "\"} +*} + text {* For an example, see HOL/ex/CodeRandom.thy (more examples coming soon). *}