# HG changeset patch # User huffman # Date 1290813068 28800 # Node ID 6f65843e78f3fe774eb44aebd6c2c5ede896a46b # Parent a292fc5157f886e153251c652d6dabf47793069e remove case combinator for fixrec match type diff -r a292fc5157f8 -r 6f65843e78f3 src/HOLCF/Fixrec.thy --- a/src/HOLCF/Fixrec.thy Fri Nov 26 14:10:34 2010 -0800 +++ b/src/HOLCF/Fixrec.thy Fri Nov 26 15:11:08 2010 -0800 @@ -26,10 +26,6 @@ succeed :: "'a \ 'a match" where "succeed = (\ x. Abs_match (sinr\(up\x)))" -definition - match_case :: "'b \ ('a \ 'b) \ 'a match \ 'b::pcpo" where - "match_case = (\ f r m. sscase\(\ x. f)\(fup\r)\(Rep_match m))" - lemma matchE [case_names bottom fail succeed, cases type: match]: "\p = \ \ Q; p = fail \ Q; \x. p = succeed\x \ Q\ \ Q" unfolding fail_def succeed_def @@ -52,40 +48,31 @@ "succeed\x \ fail" "fail \ succeed\x" by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject) -lemma match_case_simps [simp]: - "match_case\f\r\\ = \" - "match_case\f\r\fail = f" - "match_case\f\r\(succeed\x) = r\x" -by (simp_all add: succeed_def fail_def match_case_def cont_Rep_match - cont2cont_LAM - cont_Abs_match Abs_match_inverse Rep_match_strict) - -translations - "case m of XCONST fail \ t1 | XCONST succeed\x \ t2" - == "CONST match_case\t1\(\ x. t2)\m" - subsubsection {* Run operator *} definition run :: "'a match \ 'a::pcpo" where - "run = match_case\\\ID" + "run = (\ m. sscase\\\(fup\ID)\(Rep_match m))" text {* rewrite rules for run *} lemma run_strict [simp]: "run\\ = \" -by (simp add: run_def) +unfolding run_def +by (simp add: cont_Rep_match Rep_match_strict) lemma run_fail [simp]: "run\fail = \" -by (simp add: run_def) +unfolding run_def fail_def +by (simp add: cont_Rep_match Abs_match_inverse) lemma run_succeed [simp]: "run\(succeed\x) = x" -by (simp add: run_def) +unfolding run_def succeed_def +by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse) subsubsection {* Monad plus operator *} definition mplus :: "'a match \ 'a match \ 'a match" where - "mplus = (\ m1 m2. case m1 of fail \ m2 | succeed\x \ m1)" + "mplus = (\ m1 m2. sscase\(\ _. m2)\(\ _. m1)\(Rep_match m1))" abbreviation mplus_syn :: "['a match, 'a match] \ 'a match" (infixr "+++" 65) where @@ -93,14 +80,19 @@ text {* rewrite rules for mplus *} +lemmas cont2cont_Rep_match = cont_Rep_match [THEN cont_compose] + lemma mplus_strict [simp]: "\ +++ m = \" -by (simp add: mplus_def) +unfolding mplus_def +by (simp add: cont2cont_Rep_match Rep_match_strict) lemma mplus_fail [simp]: "fail +++ m = m" -by (simp add: mplus_def) +unfolding mplus_def fail_def +by (simp add: cont2cont_Rep_match Abs_match_inverse) lemma mplus_succeed [simp]: "succeed\x +++ m = succeed\x" -by (simp add: mplus_def) +unfolding mplus_def succeed_def +by (simp add: cont2cont_Rep_match cont_Abs_match Abs_match_inverse) lemma mplus_fail2 [simp]: "m +++ fail = m" by (cases m, simp_all) diff -r a292fc5157f8 -r 6f65843e78f3 src/HOLCF/ex/Pattern_Match.thy --- a/src/HOLCF/ex/Pattern_Match.thy Fri Nov 26 14:10:34 2010 -0800 +++ b/src/HOLCF/ex/Pattern_Match.thy Fri Nov 26 15:11:08 2010 -0800 @@ -53,11 +53,24 @@ lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3 +subsection {* Bind operator for match monad *} + +definition match_bind :: "'a match \ ('a \ 'b match) \ 'b match" where + "match_bind = (\ m k. sscase\(\ _. fail)\(fup\k)\(Rep_match m))" + +lemma match_bind_simps [simp]: + "match_bind\\\k = \" + "match_bind\fail\k = fail" + "match_bind\(succeed\x)\k = k\x" +unfolding match_bind_def fail_def succeed_def +by (simp_all add: cont2cont_Rep_match cont_Abs_match + Rep_match_strict Abs_match_inverse) + subsection {* Case branch combinator *} definition branch :: "('a \ 'b match) \ ('b \ 'c) \ ('a \ 'c match)" where - "branch p \ \ r x. match_case\fail\(\ y. succeed\(r\y))\(p\x)" + "branch p \ \ r x. match_bind\(p\x)\(\ y. succeed\(r\y))" lemma branch_simps: "p\x = \ \ branch p\r\x = \" @@ -72,7 +85,7 @@ definition cases :: "'a match \ 'a::pcpo" where - "cases = match_case\\\ID" + "cases = Fixrec.run" text {* rewrite rules for cases *} @@ -165,7 +178,7 @@ definition cpair_pat :: "('a, 'c) pat \ ('b, 'd) pat \ ('a \ 'b, 'c \ 'd) pat" where "cpair_pat p1 p2 = (\(x, y). - match_case\fail\(\ a. match_case\fail\(\ b. succeed\(a, b))\(p2\y))\(p1\x))" + match_bind\(p1\x)\(\ a. match_bind\(p2\y)\(\ b. succeed\(a, b))))" definition spair_pat :: @@ -313,7 +326,7 @@ definition as_pat :: "('a \ 'b match) \ 'a \ ('a \ 'b) match" where - "as_pat p = (\ x. match_case\fail\(\ a. succeed\(x, a))\(p\x))" + "as_pat p = (\ x. match_bind\(p\x)\(\ a. succeed\(x, a)))" definition lazy_pat :: "('a \ 'b::pcpo match) \ ('a \ 'b match)" where @@ -544,7 +557,7 @@ val (fun1, fun2, taken) = pat_lhs (pat, args); val defs = @{thm branch_def} :: pat_defs; val goal = mk_trp (mk_strict fun1); - val rules = @{thms match_case_simps} @ case_rews; + val rules = @{thms match_bind_simps} @ case_rews; val tacs = [simp_tac (beta_ss addsimps rules) 1]; in prove thy defs goal (K tacs) end; fun pat_apps (i, (pat, (con, args))) = @@ -559,7 +572,7 @@ val concl = mk_trp (mk_eq (fun1 ` con_app, rhs)); val goal = Logic.list_implies (assms, concl); val defs = @{thm branch_def} :: pat_defs; - val rules = @{thms match_case_simps} @ case_rews; + val rules = @{thms match_bind_simps} @ case_rews; val tacs = [asm_simp_tac (beta_ss addsimps rules) 1]; in prove thy defs goal (K tacs) end; in map_index pat_app spec end;