remove case combinator for fixrec match type
authorhuffman
Fri, 26 Nov 2010 15:11:08 -0800
changeset 40735 6f65843e78f3
parent 40734 a292fc5157f8
child 40736 72857de90621
remove case combinator for fixrec match type
src/HOLCF/Fixrec.thy
src/HOLCF/ex/Pattern_Match.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 \<rightarrow> 'a match" where
   "succeed = (\<Lambda> x. Abs_match (sinr\<cdot>(up\<cdot>x)))"
 
-definition
-  match_case :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a match \<rightarrow> 'b::pcpo" where
-  "match_case = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_match m))"
-
 lemma matchE [case_names bottom fail succeed, cases type: match]:
   "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = succeed\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
 unfolding fail_def succeed_def
@@ -52,40 +48,31 @@
   "succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x"
 by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject)
 
-lemma match_case_simps [simp]:
-  "match_case\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>"
-  "match_case\<cdot>f\<cdot>r\<cdot>fail = f"
-  "match_case\<cdot>f\<cdot>r\<cdot>(succeed\<cdot>x) = r\<cdot>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 \<Rightarrow> t1 | XCONST succeed\<cdot>x \<Rightarrow> t2"
-    == "CONST match_case\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
-
 subsubsection {* Run operator *}
 
 definition
   run :: "'a match \<rightarrow> 'a::pcpo" where
-  "run = match_case\<cdot>\<bottom>\<cdot>ID"
+  "run = (\<Lambda> m. sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)\<cdot>(Rep_match m))"
 
 text {* rewrite rules for run *}
 
 lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
-by (simp add: run_def)
+unfolding run_def
+by (simp add: cont_Rep_match Rep_match_strict)
 
 lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
-by (simp add: run_def)
+unfolding run_def fail_def
+by (simp add: cont_Rep_match Abs_match_inverse)
 
 lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>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 \<rightarrow> 'a match \<rightarrow> 'a match" where
-  "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | succeed\<cdot>x \<Rightarrow> m1)"
+  "mplus = (\<Lambda> m1 m2. sscase\<cdot>(\<Lambda> _. m2)\<cdot>(\<Lambda> _. m1)\<cdot>(Rep_match m1))"
 
 abbreviation
   mplus_syn :: "['a match, 'a match] \<Rightarrow> '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]: "\<bottom> +++ m = \<bottom>"
-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\<cdot>x +++ m = succeed\<cdot>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)
--- 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 \<rightarrow> ('a \<rightarrow> 'b match) \<rightarrow> 'b match" where
+  "match_bind = (\<Lambda> m k. sscase\<cdot>(\<Lambda> _. fail)\<cdot>(fup\<cdot>k)\<cdot>(Rep_match m))"
+
+lemma match_bind_simps [simp]:
+  "match_bind\<cdot>\<bottom>\<cdot>k = \<bottom>"
+  "match_bind\<cdot>fail\<cdot>k = fail"
+  "match_bind\<cdot>(succeed\<cdot>x)\<cdot>k = k\<cdot>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 \<rightarrow> 'b match) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c match)" where
-  "branch p \<equiv> \<Lambda> r x. match_case\<cdot>fail\<cdot>(\<Lambda> y. succeed\<cdot>(r\<cdot>y))\<cdot>(p\<cdot>x)"
+  "branch p \<equiv> \<Lambda> r x. match_bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> y. succeed\<cdot>(r\<cdot>y))"
 
 lemma branch_simps:
   "p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"
@@ -72,7 +85,7 @@
 
 definition
   cases :: "'a match \<rightarrow> 'a::pcpo" where
-  "cases = match_case\<cdot>\<bottom>\<cdot>ID"
+  "cases = Fixrec.run"
 
 text {* rewrite rules for cases *}
 
@@ -165,7 +178,7 @@
 definition
   cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
   "cpair_pat p1 p2 = (\<Lambda>(x, y).
-    match_case\<cdot>fail\<cdot>(\<Lambda> a. match_case\<cdot>fail\<cdot>(\<Lambda> b. succeed\<cdot>(a, b))\<cdot>(p2\<cdot>y))\<cdot>(p1\<cdot>x))"
+    match_bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. match_bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. succeed\<cdot>(a, b))))"
 
 definition
   spair_pat ::
@@ -313,7 +326,7 @@
 
 definition
   as_pat :: "('a \<rightarrow> 'b match) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) match" where
-  "as_pat p = (\<Lambda> x. match_case\<cdot>fail\<cdot>(\<Lambda> a. succeed\<cdot>(x, a))\<cdot>(p\<cdot>x))"
+  "as_pat p = (\<Lambda> x. match_bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. succeed\<cdot>(x, a)))"
 
 definition
   lazy_pat :: "('a \<rightarrow> 'b::pcpo match) \<Rightarrow> ('a \<rightarrow> '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;