rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
authorhuffman
Mon, 24 May 2010 09:32:52 -0700
changeset 3709200f13d3ad474
parent 37091 1535aa1c943a
child 37093 e67760c1b851
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
src/HOLCF/Fixrec.thy
src/HOLCF/Tools/Domain/domain_constructors.ML
src/HOLCF/Tools/fixrec.ML
src/HOLCF/Tools/holcf_library.ML
     1.1 --- a/src/HOLCF/Fixrec.thy	Mon May 24 13:48:57 2010 +0200
     1.2 +++ b/src/HOLCF/Fixrec.thy	Mon May 24 09:32:52 2010 -0700
     1.3 @@ -11,65 +11,65 @@
     1.4    ("Tools/fixrec.ML")
     1.5  begin
     1.6  
     1.7 -subsection {* Maybe monad type *}
     1.8 +subsection {* Pattern-match monad *}
     1.9  
    1.10  default_sort cpo
    1.11  
    1.12 -pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set"
    1.13 +pcpodef (open) 'a match = "UNIV::(one ++ 'a u) set"
    1.14  by simp_all
    1.15  
    1.16  definition
    1.17 -  fail :: "'a maybe" where
    1.18 -  "fail = Abs_maybe (sinl\<cdot>ONE)"
    1.19 +  fail :: "'a match" where
    1.20 +  "fail = Abs_match (sinl\<cdot>ONE)"
    1.21  
    1.22  definition
    1.23 -  return :: "'a \<rightarrow> 'a maybe" where
    1.24 -  "return = (\<Lambda> x. Abs_maybe (sinr\<cdot>(up\<cdot>x)))"
    1.25 +  succeed :: "'a \<rightarrow> 'a match" where
    1.26 +  "succeed = (\<Lambda> x. Abs_match (sinr\<cdot>(up\<cdot>x)))"
    1.27  
    1.28  definition
    1.29 -  maybe_when :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a maybe \<rightarrow> 'b::pcpo" where
    1.30 -  "maybe_when = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_maybe m))"
    1.31 +  match_case :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a match \<rightarrow> 'b::pcpo" where
    1.32 +  "match_case = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_match m))"
    1.33  
    1.34 -lemma maybeE:
    1.35 -  "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = return\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
    1.36 -apply (unfold fail_def return_def)
    1.37 +lemma matchE [case_names bottom fail succeed, cases type: match]:
    1.38 +  "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = succeed\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
    1.39 +unfolding fail_def succeed_def
    1.40  apply (cases p, rename_tac r)
    1.41 -apply (rule_tac p=r in ssumE, simp add: Abs_maybe_strict)
    1.42 +apply (rule_tac p=r in ssumE, simp add: Abs_match_strict)
    1.43  apply (rule_tac p=x in oneE, simp, simp)
    1.44 -apply (rule_tac p=y in upE, simp, simp add: cont_Abs_maybe)
    1.45 +apply (rule_tac p=y in upE, simp, simp add: cont_Abs_match)
    1.46  done
    1.47  
    1.48 -lemma return_defined [simp]: "return\<cdot>x \<noteq> \<bottom>"
    1.49 -by (simp add: return_def cont_Abs_maybe Abs_maybe_defined)
    1.50 +lemma succeed_defined [simp]: "succeed\<cdot>x \<noteq> \<bottom>"
    1.51 +by (simp add: succeed_def cont_Abs_match Abs_match_defined)
    1.52  
    1.53  lemma fail_defined [simp]: "fail \<noteq> \<bottom>"
    1.54 -by (simp add: fail_def Abs_maybe_defined)
    1.55 +by (simp add: fail_def Abs_match_defined)
    1.56  
    1.57 -lemma return_eq [simp]: "(return\<cdot>x = return\<cdot>y) = (x = y)"
    1.58 -by (simp add: return_def cont_Abs_maybe Abs_maybe_inject)
    1.59 +lemma succeed_eq [simp]: "(succeed\<cdot>x = succeed\<cdot>y) = (x = y)"
    1.60 +by (simp add: succeed_def cont_Abs_match Abs_match_inject)
    1.61  
    1.62 -lemma return_neq_fail [simp]:
    1.63 -  "return\<cdot>x \<noteq> fail" "fail \<noteq> return\<cdot>x"
    1.64 -by (simp_all add: return_def fail_def cont_Abs_maybe Abs_maybe_inject)
    1.65 +lemma succeed_neq_fail [simp]:
    1.66 +  "succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x"
    1.67 +by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject)
    1.68  
    1.69 -lemma maybe_when_rews [simp]:
    1.70 -  "maybe_when\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>"
    1.71 -  "maybe_when\<cdot>f\<cdot>r\<cdot>fail = f"
    1.72 -  "maybe_when\<cdot>f\<cdot>r\<cdot>(return\<cdot>x) = r\<cdot>x"
    1.73 -by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe
    1.74 +lemma match_case_simps [simp]:
    1.75 +  "match_case\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>"
    1.76 +  "match_case\<cdot>f\<cdot>r\<cdot>fail = f"
    1.77 +  "match_case\<cdot>f\<cdot>r\<cdot>(succeed\<cdot>x) = r\<cdot>x"
    1.78 +by (simp_all add: succeed_def fail_def match_case_def cont_Rep_match
    1.79                    cont2cont_LAM
    1.80 -                  cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict)
    1.81 +                  cont_Abs_match Abs_match_inverse Rep_match_strict)
    1.82  
    1.83  translations
    1.84 -  "case m of XCONST fail \<Rightarrow> t1 | XCONST return\<cdot>x \<Rightarrow> t2"
    1.85 -    == "CONST maybe_when\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
    1.86 +  "case m of XCONST fail \<Rightarrow> t1 | XCONST succeed\<cdot>x \<Rightarrow> t2"
    1.87 +    == "CONST match_case\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
    1.88  
    1.89  
    1.90  subsubsection {* Run operator *}
    1.91  
    1.92  definition
    1.93 -  run :: "'a maybe \<rightarrow> 'a::pcpo" where
    1.94 -  "run = maybe_when\<cdot>\<bottom>\<cdot>ID"
    1.95 +  run :: "'a match \<rightarrow> 'a::pcpo" where
    1.96 +  "run = match_case\<cdot>\<bottom>\<cdot>ID"
    1.97  
    1.98  text {* rewrite rules for run *}
    1.99  
   1.100 @@ -79,17 +79,17 @@
   1.101  lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
   1.102  by (simp add: run_def)
   1.103  
   1.104 -lemma run_return [simp]: "run\<cdot>(return\<cdot>x) = x"
   1.105 +lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>x) = x"
   1.106  by (simp add: run_def)
   1.107  
   1.108  subsubsection {* Monad plus operator *}
   1.109  
   1.110  definition
   1.111 -  mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe" where
   1.112 -  "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | return\<cdot>x \<Rightarrow> m1)"
   1.113 +  mplus :: "'a match \<rightarrow> 'a match \<rightarrow> 'a match" where
   1.114 +  "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | succeed\<cdot>x \<Rightarrow> m1)"
   1.115  
   1.116  abbreviation
   1.117 -  mplus_syn :: "['a maybe, 'a maybe] \<Rightarrow> 'a maybe"  (infixr "+++" 65)  where
   1.118 +  mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match"  (infixr "+++" 65)  where
   1.119    "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
   1.120  
   1.121  text {* rewrite rules for mplus *}
   1.122 @@ -100,23 +100,23 @@
   1.123  lemma mplus_fail [simp]: "fail +++ m = m"
   1.124  by (simp add: mplus_def)
   1.125  
   1.126 -lemma mplus_return [simp]: "return\<cdot>x +++ m = return\<cdot>x"
   1.127 +lemma mplus_succeed [simp]: "succeed\<cdot>x +++ m = succeed\<cdot>x"
   1.128  by (simp add: mplus_def)
   1.129  
   1.130  lemma mplus_fail2 [simp]: "m +++ fail = m"
   1.131 -by (rule_tac p=m in maybeE, simp_all)
   1.132 +by (cases m, simp_all)
   1.133  
   1.134  lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
   1.135 -by (rule_tac p=x in maybeE, simp_all)
   1.136 +by (cases x, simp_all)
   1.137  
   1.138  subsubsection {* Fatbar combinator *}
   1.139  
   1.140  definition
   1.141 -  fatbar :: "('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe)" where
   1.142 +  fatbar :: "('a \<rightarrow> 'b match) \<rightarrow> ('a \<rightarrow> 'b match) \<rightarrow> ('a \<rightarrow> 'b match)" where
   1.143    "fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)"
   1.144  
   1.145  abbreviation
   1.146 -  fatbar_syn :: "['a \<rightarrow> 'b maybe, 'a \<rightarrow> 'b maybe] \<Rightarrow> 'a \<rightarrow> 'b maybe" (infixr "\<parallel>" 60)  where
   1.147 +  fatbar_syn :: "['a \<rightarrow> 'b match, 'a \<rightarrow> 'b match] \<Rightarrow> 'a \<rightarrow> 'b match" (infixr "\<parallel>" 60)  where
   1.148    "m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2"
   1.149  
   1.150  lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>"
   1.151 @@ -125,7 +125,7 @@
   1.152  lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x"
   1.153  by (simp add: fatbar_def)
   1.154  
   1.155 -lemma fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = return\<cdot>y"
   1.156 +lemma fatbar3: "m\<cdot>x = succeed\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = succeed\<cdot>y"
   1.157  by (simp add: fatbar_def)
   1.158  
   1.159  lemmas fatbar_simps = fatbar1 fatbar2 fatbar3
   1.160 @@ -136,7 +136,7 @@
   1.161  lemma run_fatbar2: "m\<cdot>x = fail \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = run\<cdot>(ms\<cdot>x)"
   1.162  by (simp add: fatbar_def)
   1.163  
   1.164 -lemma run_fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y"
   1.165 +lemma run_fatbar3: "m\<cdot>x = succeed\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y"
   1.166  by (simp add: fatbar_def)
   1.167  
   1.168  lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3
   1.169 @@ -144,23 +144,23 @@
   1.170  subsection {* Case branch combinator *}
   1.171  
   1.172  definition
   1.173 -  branch :: "('a \<rightarrow> 'b maybe) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c maybe)" where
   1.174 -  "branch p \<equiv> \<Lambda> r x. maybe_when\<cdot>fail\<cdot>(\<Lambda> y. return\<cdot>(r\<cdot>y))\<cdot>(p\<cdot>x)"
   1.175 +  branch :: "('a \<rightarrow> 'b match) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c match)" where
   1.176 +  "branch p \<equiv> \<Lambda> r x. match_case\<cdot>fail\<cdot>(\<Lambda> y. succeed\<cdot>(r\<cdot>y))\<cdot>(p\<cdot>x)"
   1.177  
   1.178 -lemma branch_rews:
   1.179 +lemma branch_simps:
   1.180    "p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"
   1.181    "p\<cdot>x = fail \<Longrightarrow> branch p\<cdot>r\<cdot>x = fail"
   1.182 -  "p\<cdot>x = return\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>y)"
   1.183 +  "p\<cdot>x = succeed\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>y)"
   1.184  by (simp_all add: branch_def)
   1.185  
   1.186 -lemma branch_return [simp]: "branch return\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>x)"
   1.187 +lemma branch_succeed [simp]: "branch succeed\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>x)"
   1.188  by (simp add: branch_def)
   1.189  
   1.190  subsubsection {* Cases operator *}
   1.191  
   1.192  definition
   1.193 -  cases :: "'a maybe \<rightarrow> 'a::pcpo" where
   1.194 -  "cases = maybe_when\<cdot>\<bottom>\<cdot>ID"
   1.195 +  cases :: "'a match \<rightarrow> 'a::pcpo" where
   1.196 +  "cases = match_case\<cdot>\<bottom>\<cdot>ID"
   1.197  
   1.198  text {* rewrite rules for cases *}
   1.199  
   1.200 @@ -170,7 +170,7 @@
   1.201  lemma cases_fail [simp]: "cases\<cdot>fail = \<bottom>"
   1.202  by (simp add: cases_def)
   1.203  
   1.204 -lemma cases_return [simp]: "cases\<cdot>(return\<cdot>x) = x"
   1.205 +lemma cases_succeed [simp]: "cases\<cdot>(succeed\<cdot>x) = x"
   1.206  by (simp add: cases_def)
   1.207  
   1.208  subsection {* Case syntax *}
   1.209 @@ -204,9 +204,9 @@
   1.210    "_variable _noargs r" => "CONST unit_when\<cdot>r"
   1.211  
   1.212  parse_translation {*
   1.213 -(* rewrite (_pat x) => (return) *)
   1.214 +(* rewrite (_pat x) => (succeed) *)
   1.215  (* rewrite (_variable x t) => (Abs_CFun (%x. t)) *)
   1.216 - [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.return}),
   1.217 + [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.succeed}),
   1.218    mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_CFun})];
   1.219  *}
   1.220  
   1.221 @@ -243,17 +243,17 @@
   1.222  *}
   1.223  
   1.224  translations
   1.225 -  "x" <= "_match (CONST Fixrec.return) (_variable x)"
   1.226 +  "x" <= "_match (CONST Fixrec.succeed) (_variable x)"
   1.227  
   1.228  
   1.229  subsection {* Pattern combinators for data constructors *}
   1.230  
   1.231 -types ('a, 'b) pat = "'a \<rightarrow> 'b maybe"
   1.232 +types ('a, 'b) pat = "'a \<rightarrow> 'b match"
   1.233  
   1.234  definition
   1.235    cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
   1.236    "cpair_pat p1 p2 = (\<Lambda>(x, y).
   1.237 -    maybe_when\<cdot>fail\<cdot>(\<Lambda> a. maybe_when\<cdot>fail\<cdot>(\<Lambda> b. return\<cdot>(a, b))\<cdot>(p2\<cdot>y))\<cdot>(p1\<cdot>x))"
   1.238 +    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))"
   1.239  
   1.240  definition
   1.241    spair_pat ::
   1.242 @@ -274,15 +274,15 @@
   1.243  
   1.244  definition
   1.245    TT_pat :: "(tr, unit) pat" where
   1.246 -  "TT_pat = (\<Lambda> b. If b then return\<cdot>() else fail fi)"
   1.247 +  "TT_pat = (\<Lambda> b. If b then succeed\<cdot>() else fail fi)"
   1.248  
   1.249  definition
   1.250    FF_pat :: "(tr, unit) pat" where
   1.251 -  "FF_pat = (\<Lambda> b. If b then fail else return\<cdot>() fi)"
   1.252 +  "FF_pat = (\<Lambda> b. If b then fail else succeed\<cdot>() fi)"
   1.253  
   1.254  definition
   1.255    ONE_pat :: "(one, unit) pat" where
   1.256 -  "ONE_pat = (\<Lambda> ONE. return\<cdot>())"
   1.257 +  "ONE_pat = (\<Lambda> ONE. succeed\<cdot>())"
   1.258  
   1.259  text {* Parse translations (patterns) *}
   1.260  translations
   1.261 @@ -331,21 +331,21 @@
   1.262  lemma cpair_pat1:
   1.263    "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = \<bottom>"
   1.264  apply (simp add: branch_def cpair_pat_def)
   1.265 -apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
   1.266 +apply (cases "p\<cdot>x", simp_all)
   1.267  done
   1.268  
   1.269  lemma cpair_pat2:
   1.270    "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = fail"
   1.271  apply (simp add: branch_def cpair_pat_def)
   1.272 -apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
   1.273 +apply (cases "p\<cdot>x", simp_all)
   1.274  done
   1.275  
   1.276  lemma cpair_pat3:
   1.277 -  "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow>
   1.278 +  "branch p\<cdot>r\<cdot>x = succeed\<cdot>s \<Longrightarrow>
   1.279     branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = branch q\<cdot>s\<cdot>y"
   1.280  apply (simp add: branch_def cpair_pat_def)
   1.281 -apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
   1.282 -apply (rule_tac p="q\<cdot>y" in maybeE, simp_all)
   1.283 +apply (cases "p\<cdot>x", simp_all)
   1.284 +apply (cases "q\<cdot>y", simp_all)
   1.285  done
   1.286  
   1.287  lemmas cpair_pat [simp] =
   1.288 @@ -377,35 +377,35 @@
   1.289  
   1.290  lemma TT_pat [simp]:
   1.291    "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   1.292 -  "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = return\<cdot>r"
   1.293 +  "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = succeed\<cdot>r"
   1.294    "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail"
   1.295  by (simp_all add: branch_def TT_pat_def)
   1.296  
   1.297  lemma FF_pat [simp]:
   1.298    "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   1.299    "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail"
   1.300 -  "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = return\<cdot>r"
   1.301 +  "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = succeed\<cdot>r"
   1.302  by (simp_all add: branch_def FF_pat_def)
   1.303  
   1.304  lemma ONE_pat [simp]:
   1.305    "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   1.306 -  "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = return\<cdot>r"
   1.307 +  "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = succeed\<cdot>r"
   1.308  by (simp_all add: branch_def ONE_pat_def)
   1.309  
   1.310  
   1.311  subsection {* Wildcards, as-patterns, and lazy patterns *}
   1.312  
   1.313  definition
   1.314 -  wild_pat :: "'a \<rightarrow> unit maybe" where
   1.315 -  "wild_pat = (\<Lambda> x. return\<cdot>())"
   1.316 +  wild_pat :: "'a \<rightarrow> unit match" where
   1.317 +  "wild_pat = (\<Lambda> x. succeed\<cdot>())"
   1.318  
   1.319  definition
   1.320 -  as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
   1.321 -  "as_pat p = (\<Lambda> x. maybe_when\<cdot>fail\<cdot>(\<Lambda> a. return\<cdot>(x, a))\<cdot>(p\<cdot>x))"
   1.322 +  as_pat :: "('a \<rightarrow> 'b match) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) match" where
   1.323 +  "as_pat p = (\<Lambda> x. match_case\<cdot>fail\<cdot>(\<Lambda> a. succeed\<cdot>(x, a))\<cdot>(p\<cdot>x))"
   1.324  
   1.325  definition
   1.326 -  lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
   1.327 -  "lazy_pat p = (\<Lambda> x. return\<cdot>(cases\<cdot>(p\<cdot>x)))"
   1.328 +  lazy_pat :: "('a \<rightarrow> 'b::pcpo match) \<Rightarrow> ('a \<rightarrow> 'b match)" where
   1.329 +  "lazy_pat p = (\<Lambda> x. succeed\<cdot>(cases\<cdot>(p\<cdot>x)))"
   1.330  
   1.331  text {* Parse translations (patterns) *}
   1.332  translations
   1.333 @@ -419,21 +419,21 @@
   1.334  translations
   1.335    "_" <= "_match (CONST wild_pat) _noargs"
   1.336  
   1.337 -lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = return\<cdot>r"
   1.338 +lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = succeed\<cdot>r"
   1.339  by (simp add: branch_def wild_pat_def)
   1.340  
   1.341  lemma as_pat [simp]:
   1.342    "branch (as_pat p)\<cdot>(csplit\<cdot>r)\<cdot>x = branch p\<cdot>(r\<cdot>x)\<cdot>x"
   1.343  apply (simp add: branch_def as_pat_def)
   1.344 -apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
   1.345 +apply (cases "p\<cdot>x", simp_all)
   1.346  done
   1.347  
   1.348  lemma lazy_pat [simp]:
   1.349 -  "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)"
   1.350 -  "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)"
   1.351 -  "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>s"
   1.352 +  "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>\<bottom>)"
   1.353 +  "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>\<bottom>)"
   1.354 +  "branch p\<cdot>r\<cdot>x = succeed\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>s"
   1.355  apply (simp_all add: branch_def lazy_pat_def)
   1.356 -apply (rule_tac [!] p="p\<cdot>x" in maybeE, simp_all)
   1.357 +apply (cases "p\<cdot>x", simp_all)+
   1.358  done
   1.359  
   1.360  
   1.361 @@ -442,47 +442,47 @@
   1.362  default_sort pcpo
   1.363  
   1.364  definition
   1.365 -  match_UU :: "'a \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
   1.366 +  match_UU :: "'a \<rightarrow> 'c match \<rightarrow> 'c match"
   1.367  where
   1.368    "match_UU = strictify\<cdot>(\<Lambda> x k. fail)"
   1.369  
   1.370  definition
   1.371 -  match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   1.372 +  match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match"
   1.373  where
   1.374    "match_cpair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)"
   1.375  
   1.376  definition
   1.377 -  match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   1.378 +  match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match"
   1.379  where
   1.380    "match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)"
   1.381  
   1.382  definition
   1.383 -  match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   1.384 +  match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match"
   1.385  where
   1.386    "match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)"
   1.387  
   1.388  definition
   1.389 -  match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   1.390 +  match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c match) \<rightarrow> 'c match"
   1.391  where
   1.392    "match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)"
   1.393  
   1.394  definition
   1.395 -  match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   1.396 +  match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match"
   1.397  where
   1.398    "match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)"
   1.399  
   1.400  definition
   1.401 -  match_ONE :: "one \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
   1.402 +  match_ONE :: "one \<rightarrow> 'c match \<rightarrow> 'c match"
   1.403  where
   1.404    "match_ONE = (\<Lambda> ONE k. k)"
   1.405  
   1.406  definition
   1.407 -  match_TT :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
   1.408 +  match_TT :: "tr \<rightarrow> 'c match \<rightarrow> 'c match"
   1.409  where
   1.410    "match_TT = (\<Lambda> x k. If x then k else fail fi)"
   1.411   
   1.412  definition
   1.413 -  match_FF :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
   1.414 +  match_FF :: "tr \<rightarrow> 'c match \<rightarrow> 'c match"
   1.415  where
   1.416    "match_FF = (\<Lambda> x k. If x then fail else k fi)"
   1.417  
   1.418 @@ -584,6 +584,6 @@
   1.419        (@{const_name UU}, @{const_name match_UU}) ]
   1.420  *}
   1.421  
   1.422 -hide_const (open) return fail run cases
   1.423 +hide_const (open) succeed fail run cases
   1.424  
   1.425  end
     2.1 --- a/src/HOLCF/Tools/Domain/domain_constructors.ML	Mon May 24 13:48:57 2010 +0200
     2.2 +++ b/src/HOLCF/Tools/Domain/domain_constructors.ML	Mon May 24 09:32:52 2010 -0700
     2.3 @@ -850,7 +850,7 @@
     2.4        in
     2.5          pat_const $ p1 $ p2
     2.6        end;
     2.7 -    fun mk_tuple_pat [] = return_const HOLogic.unitT
     2.8 +    fun mk_tuple_pat [] = succeed_const HOLogic.unitT
     2.9        | mk_tuple_pat ps = foldr1 mk_pair_pat ps;
    2.10      fun branch_const (T,U,V) = 
    2.11        Const (@{const_name branch},
    2.12 @@ -951,7 +951,7 @@
    2.13            val (fun1, fun2, taken) = pat_lhs (pat, args);
    2.14            val defs = @{thm branch_def} :: pat_defs;
    2.15            val goal = mk_trp (mk_strict fun1);
    2.16 -          val rules = @{thms maybe_when_rews} @ case_rews;
    2.17 +          val rules = @{thms match_case_simps} @ case_rews;
    2.18            val tacs = [simp_tac (beta_ss addsimps rules) 1];
    2.19          in prove thy defs goal (K tacs) end;
    2.20        fun pat_apps (i, (pat, (con, args))) =
    2.21 @@ -966,7 +966,7 @@
    2.22                val concl = mk_trp (mk_eq (fun1 ` con_app, rhs));
    2.23                val goal = Logic.list_implies (assms, concl);
    2.24                val defs = @{thm branch_def} :: pat_defs;
    2.25 -              val rules = @{thms maybe_when_rews} @ case_rews;
    2.26 +              val rules = @{thms match_case_simps} @ case_rews;
    2.27                val tacs = [asm_simp_tac (beta_ss addsimps rules) 1];
    2.28              in prove thy defs goal (K tacs) end;
    2.29          in map_index pat_app spec end;
     3.1 --- a/src/HOLCF/Tools/fixrec.ML	Mon May 24 13:48:57 2010 +0200
     3.2 +++ b/src/HOLCF/Tools/fixrec.ML	Mon May 24 09:32:52 2010 -0700
     3.3 @@ -85,7 +85,7 @@
     3.4    in
     3.5      case t of
     3.6        Const(@{const_name Rep_CFun}, _) $
     3.7 -        Const(@{const_name Fixrec.return}, _) $ u => u
     3.8 +        Const(@{const_name Fixrec.succeed}, _) $ u => u
     3.9      | _ => run ` t
    3.10    end;
    3.11  
    3.12 @@ -267,7 +267,7 @@
    3.13    let
    3.14      val (lhs,rhs) = dest_eqs (Logic.strip_imp_concl (strip_alls eq));
    3.15    in
    3.16 -    building match_name lhs (mk_return rhs) [] (taken_names eq)
    3.17 +    building match_name lhs (mk_succeed rhs) [] (taken_names eq)
    3.18    end;
    3.19  
    3.20  (* returns the sum (using +++) of the terms in ms *)
     4.1 --- a/src/HOLCF/Tools/holcf_library.ML	Mon May 24 13:48:57 2010 +0200
     4.2 +++ b/src/HOLCF/Tools/holcf_library.ML	Mon May 24 09:32:52 2010 -0700
     4.3 @@ -252,15 +252,15 @@
     4.4  
     4.5  (*** pattern match monad type ***)
     4.6  
     4.7 -fun mk_matchT T = Type (@{type_name "maybe"}, [T]);
     4.8 +fun mk_matchT T = Type (@{type_name "match"}, [T]);
     4.9  
    4.10 -fun dest_matchT (Type(@{type_name "maybe"}, [T])) = T
    4.11 +fun dest_matchT (Type(@{type_name "match"}, [T])) = T
    4.12    | dest_matchT T = raise TYPE ("dest_matchT", [T], []);
    4.13  
    4.14  fun mk_fail T = Const (@{const_name "Fixrec.fail"}, mk_matchT T);
    4.15  
    4.16 -fun return_const T = Const (@{const_name "Fixrec.return"}, T ->> mk_matchT T);
    4.17 -fun mk_return t = return_const (fastype_of t) ` t;
    4.18 +fun succeed_const T = Const (@{const_name "Fixrec.succeed"}, T ->> mk_matchT T);
    4.19 +fun mk_succeed t = succeed_const (fastype_of t) ` t;
    4.20  
    4.21  
    4.22  (*** lifted boolean type ***)