src/HOLCF/Fixrec.thy
changeset 37108 00f13d3ad474
parent 37080 a2a1c8a658ef
child 37109 e67760c1b851
equal deleted inserted replaced
37107:1535aa1c943a 37108:00f13d3ad474
     9 uses
     9 uses
    10   ("Tools/holcf_library.ML")
    10   ("Tools/holcf_library.ML")
    11   ("Tools/fixrec.ML")
    11   ("Tools/fixrec.ML")
    12 begin
    12 begin
    13 
    13 
    14 subsection {* Maybe monad type *}
    14 subsection {* Pattern-match monad *}
    15 
    15 
    16 default_sort cpo
    16 default_sort cpo
    17 
    17 
    18 pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set"
    18 pcpodef (open) 'a match = "UNIV::(one ++ 'a u) set"
    19 by simp_all
    19 by simp_all
    20 
    20 
    21 definition
    21 definition
    22   fail :: "'a maybe" where
    22   fail :: "'a match" where
    23   "fail = Abs_maybe (sinl\<cdot>ONE)"
    23   "fail = Abs_match (sinl\<cdot>ONE)"
    24 
    24 
    25 definition
    25 definition
    26   return :: "'a \<rightarrow> 'a maybe" where
    26   succeed :: "'a \<rightarrow> 'a match" where
    27   "return = (\<Lambda> x. Abs_maybe (sinr\<cdot>(up\<cdot>x)))"
    27   "succeed = (\<Lambda> x. Abs_match (sinr\<cdot>(up\<cdot>x)))"
    28 
    28 
    29 definition
    29 definition
    30   maybe_when :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a maybe \<rightarrow> 'b::pcpo" where
    30   match_case :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a match \<rightarrow> 'b::pcpo" where
    31   "maybe_when = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_maybe m))"
    31   "match_case = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_match m))"
    32 
    32 
    33 lemma maybeE:
    33 lemma matchE [case_names bottom fail succeed, cases type: match]:
    34   "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = return\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
    34   "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = succeed\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
    35 apply (unfold fail_def return_def)
    35 unfolding fail_def succeed_def
    36 apply (cases p, rename_tac r)
    36 apply (cases p, rename_tac r)
    37 apply (rule_tac p=r in ssumE, simp add: Abs_maybe_strict)
    37 apply (rule_tac p=r in ssumE, simp add: Abs_match_strict)
    38 apply (rule_tac p=x in oneE, simp, simp)
    38 apply (rule_tac p=x in oneE, simp, simp)
    39 apply (rule_tac p=y in upE, simp, simp add: cont_Abs_maybe)
    39 apply (rule_tac p=y in upE, simp, simp add: cont_Abs_match)
    40 done
    40 done
    41 
    41 
    42 lemma return_defined [simp]: "return\<cdot>x \<noteq> \<bottom>"
    42 lemma succeed_defined [simp]: "succeed\<cdot>x \<noteq> \<bottom>"
    43 by (simp add: return_def cont_Abs_maybe Abs_maybe_defined)
    43 by (simp add: succeed_def cont_Abs_match Abs_match_defined)
    44 
    44 
    45 lemma fail_defined [simp]: "fail \<noteq> \<bottom>"
    45 lemma fail_defined [simp]: "fail \<noteq> \<bottom>"
    46 by (simp add: fail_def Abs_maybe_defined)
    46 by (simp add: fail_def Abs_match_defined)
    47 
    47 
    48 lemma return_eq [simp]: "(return\<cdot>x = return\<cdot>y) = (x = y)"
    48 lemma succeed_eq [simp]: "(succeed\<cdot>x = succeed\<cdot>y) = (x = y)"
    49 by (simp add: return_def cont_Abs_maybe Abs_maybe_inject)
    49 by (simp add: succeed_def cont_Abs_match Abs_match_inject)
    50 
    50 
    51 lemma return_neq_fail [simp]:
    51 lemma succeed_neq_fail [simp]:
    52   "return\<cdot>x \<noteq> fail" "fail \<noteq> return\<cdot>x"
    52   "succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x"
    53 by (simp_all add: return_def fail_def cont_Abs_maybe Abs_maybe_inject)
    53 by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject)
    54 
    54 
    55 lemma maybe_when_rews [simp]:
    55 lemma match_case_simps [simp]:
    56   "maybe_when\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>"
    56   "match_case\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>"
    57   "maybe_when\<cdot>f\<cdot>r\<cdot>fail = f"
    57   "match_case\<cdot>f\<cdot>r\<cdot>fail = f"
    58   "maybe_when\<cdot>f\<cdot>r\<cdot>(return\<cdot>x) = r\<cdot>x"
    58   "match_case\<cdot>f\<cdot>r\<cdot>(succeed\<cdot>x) = r\<cdot>x"
    59 by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe
    59 by (simp_all add: succeed_def fail_def match_case_def cont_Rep_match
    60                   cont2cont_LAM
    60                   cont2cont_LAM
    61                   cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict)
    61                   cont_Abs_match Abs_match_inverse Rep_match_strict)
    62 
    62 
    63 translations
    63 translations
    64   "case m of XCONST fail \<Rightarrow> t1 | XCONST return\<cdot>x \<Rightarrow> t2"
    64   "case m of XCONST fail \<Rightarrow> t1 | XCONST succeed\<cdot>x \<Rightarrow> t2"
    65     == "CONST maybe_when\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
    65     == "CONST match_case\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
    66 
    66 
    67 
    67 
    68 subsubsection {* Run operator *}
    68 subsubsection {* Run operator *}
    69 
    69 
    70 definition
    70 definition
    71   run :: "'a maybe \<rightarrow> 'a::pcpo" where
    71   run :: "'a match \<rightarrow> 'a::pcpo" where
    72   "run = maybe_when\<cdot>\<bottom>\<cdot>ID"
    72   "run = match_case\<cdot>\<bottom>\<cdot>ID"
    73 
    73 
    74 text {* rewrite rules for run *}
    74 text {* rewrite rules for run *}
    75 
    75 
    76 lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
    76 lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
    77 by (simp add: run_def)
    77 by (simp add: run_def)
    78 
    78 
    79 lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
    79 lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
    80 by (simp add: run_def)
    80 by (simp add: run_def)
    81 
    81 
    82 lemma run_return [simp]: "run\<cdot>(return\<cdot>x) = x"
    82 lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>x) = x"
    83 by (simp add: run_def)
    83 by (simp add: run_def)
    84 
    84 
    85 subsubsection {* Monad plus operator *}
    85 subsubsection {* Monad plus operator *}
    86 
    86 
    87 definition
    87 definition
    88   mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe" where
    88   mplus :: "'a match \<rightarrow> 'a match \<rightarrow> 'a match" where
    89   "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | return\<cdot>x \<Rightarrow> m1)"
    89   "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | succeed\<cdot>x \<Rightarrow> m1)"
    90 
    90 
    91 abbreviation
    91 abbreviation
    92   mplus_syn :: "['a maybe, 'a maybe] \<Rightarrow> 'a maybe"  (infixr "+++" 65)  where
    92   mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match"  (infixr "+++" 65)  where
    93   "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
    93   "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
    94 
    94 
    95 text {* rewrite rules for mplus *}
    95 text {* rewrite rules for mplus *}
    96 
    96 
    97 lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
    97 lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
    98 by (simp add: mplus_def)
    98 by (simp add: mplus_def)
    99 
    99 
   100 lemma mplus_fail [simp]: "fail +++ m = m"
   100 lemma mplus_fail [simp]: "fail +++ m = m"
   101 by (simp add: mplus_def)
   101 by (simp add: mplus_def)
   102 
   102 
   103 lemma mplus_return [simp]: "return\<cdot>x +++ m = return\<cdot>x"
   103 lemma mplus_succeed [simp]: "succeed\<cdot>x +++ m = succeed\<cdot>x"
   104 by (simp add: mplus_def)
   104 by (simp add: mplus_def)
   105 
   105 
   106 lemma mplus_fail2 [simp]: "m +++ fail = m"
   106 lemma mplus_fail2 [simp]: "m +++ fail = m"
   107 by (rule_tac p=m in maybeE, simp_all)
   107 by (cases m, simp_all)
   108 
   108 
   109 lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
   109 lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
   110 by (rule_tac p=x in maybeE, simp_all)
   110 by (cases x, simp_all)
   111 
   111 
   112 subsubsection {* Fatbar combinator *}
   112 subsubsection {* Fatbar combinator *}
   113 
   113 
   114 definition
   114 definition
   115   fatbar :: "('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe)" where
   115   fatbar :: "('a \<rightarrow> 'b match) \<rightarrow> ('a \<rightarrow> 'b match) \<rightarrow> ('a \<rightarrow> 'b match)" where
   116   "fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)"
   116   "fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)"
   117 
   117 
   118 abbreviation
   118 abbreviation
   119   fatbar_syn :: "['a \<rightarrow> 'b maybe, 'a \<rightarrow> 'b maybe] \<Rightarrow> 'a \<rightarrow> 'b maybe" (infixr "\<parallel>" 60)  where
   119   fatbar_syn :: "['a \<rightarrow> 'b match, 'a \<rightarrow> 'b match] \<Rightarrow> 'a \<rightarrow> 'b match" (infixr "\<parallel>" 60)  where
   120   "m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2"
   120   "m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2"
   121 
   121 
   122 lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>"
   122 lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>"
   123 by (simp add: fatbar_def)
   123 by (simp add: fatbar_def)
   124 
   124 
   125 lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x"
   125 lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x"
   126 by (simp add: fatbar_def)
   126 by (simp add: fatbar_def)
   127 
   127 
   128 lemma fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = return\<cdot>y"
   128 lemma fatbar3: "m\<cdot>x = succeed\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = succeed\<cdot>y"
   129 by (simp add: fatbar_def)
   129 by (simp add: fatbar_def)
   130 
   130 
   131 lemmas fatbar_simps = fatbar1 fatbar2 fatbar3
   131 lemmas fatbar_simps = fatbar1 fatbar2 fatbar3
   132 
   132 
   133 lemma run_fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = \<bottom>"
   133 lemma run_fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = \<bottom>"
   134 by (simp add: fatbar_def)
   134 by (simp add: fatbar_def)
   135 
   135 
   136 lemma run_fatbar2: "m\<cdot>x = fail \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = run\<cdot>(ms\<cdot>x)"
   136 lemma run_fatbar2: "m\<cdot>x = fail \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = run\<cdot>(ms\<cdot>x)"
   137 by (simp add: fatbar_def)
   137 by (simp add: fatbar_def)
   138 
   138 
   139 lemma run_fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y"
   139 lemma run_fatbar3: "m\<cdot>x = succeed\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y"
   140 by (simp add: fatbar_def)
   140 by (simp add: fatbar_def)
   141 
   141 
   142 lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3
   142 lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3
   143 
   143 
   144 subsection {* Case branch combinator *}
   144 subsection {* Case branch combinator *}
   145 
   145 
   146 definition
   146 definition
   147   branch :: "('a \<rightarrow> 'b maybe) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c maybe)" where
   147   branch :: "('a \<rightarrow> 'b match) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c match)" where
   148   "branch p \<equiv> \<Lambda> r x. maybe_when\<cdot>fail\<cdot>(\<Lambda> y. return\<cdot>(r\<cdot>y))\<cdot>(p\<cdot>x)"
   148   "branch p \<equiv> \<Lambda> r x. match_case\<cdot>fail\<cdot>(\<Lambda> y. succeed\<cdot>(r\<cdot>y))\<cdot>(p\<cdot>x)"
   149 
   149 
   150 lemma branch_rews:
   150 lemma branch_simps:
   151   "p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"
   151   "p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"
   152   "p\<cdot>x = fail \<Longrightarrow> branch p\<cdot>r\<cdot>x = fail"
   152   "p\<cdot>x = fail \<Longrightarrow> branch p\<cdot>r\<cdot>x = fail"
   153   "p\<cdot>x = return\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>y)"
   153   "p\<cdot>x = succeed\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>y)"
   154 by (simp_all add: branch_def)
   154 by (simp_all add: branch_def)
   155 
   155 
   156 lemma branch_return [simp]: "branch return\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>x)"
   156 lemma branch_succeed [simp]: "branch succeed\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>x)"
   157 by (simp add: branch_def)
   157 by (simp add: branch_def)
   158 
   158 
   159 subsubsection {* Cases operator *}
   159 subsubsection {* Cases operator *}
   160 
   160 
   161 definition
   161 definition
   162   cases :: "'a maybe \<rightarrow> 'a::pcpo" where
   162   cases :: "'a match \<rightarrow> 'a::pcpo" where
   163   "cases = maybe_when\<cdot>\<bottom>\<cdot>ID"
   163   "cases = match_case\<cdot>\<bottom>\<cdot>ID"
   164 
   164 
   165 text {* rewrite rules for cases *}
   165 text {* rewrite rules for cases *}
   166 
   166 
   167 lemma cases_strict [simp]: "cases\<cdot>\<bottom> = \<bottom>"
   167 lemma cases_strict [simp]: "cases\<cdot>\<bottom> = \<bottom>"
   168 by (simp add: cases_def)
   168 by (simp add: cases_def)
   169 
   169 
   170 lemma cases_fail [simp]: "cases\<cdot>fail = \<bottom>"
   170 lemma cases_fail [simp]: "cases\<cdot>fail = \<bottom>"
   171 by (simp add: cases_def)
   171 by (simp add: cases_def)
   172 
   172 
   173 lemma cases_return [simp]: "cases\<cdot>(return\<cdot>x) = x"
   173 lemma cases_succeed [simp]: "cases\<cdot>(succeed\<cdot>x) = x"
   174 by (simp add: cases_def)
   174 by (simp add: cases_def)
   175 
   175 
   176 subsection {* Case syntax *}
   176 subsection {* Case syntax *}
   177 
   177 
   178 nonterminals
   178 nonterminals
   202   "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)"
   202   "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)"
   203   "_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))"
   203   "_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))"
   204   "_variable _noargs r" => "CONST unit_when\<cdot>r"
   204   "_variable _noargs r" => "CONST unit_when\<cdot>r"
   205 
   205 
   206 parse_translation {*
   206 parse_translation {*
   207 (* rewrite (_pat x) => (return) *)
   207 (* rewrite (_pat x) => (succeed) *)
   208 (* rewrite (_variable x t) => (Abs_CFun (%x. t)) *)
   208 (* rewrite (_variable x t) => (Abs_CFun (%x. t)) *)
   209  [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.return}),
   209  [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.succeed}),
   210   mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_CFun})];
   210   mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_CFun})];
   211 *}
   211 *}
   212 
   212 
   213 text {* Printing Case expressions *}
   213 text {* Printing Case expressions *}
   214 
   214 
   241 
   241 
   242   in [(@{const_syntax Rep_CFun}, Case1_tr')] end;
   242   in [(@{const_syntax Rep_CFun}, Case1_tr')] end;
   243 *}
   243 *}
   244 
   244 
   245 translations
   245 translations
   246   "x" <= "_match (CONST Fixrec.return) (_variable x)"
   246   "x" <= "_match (CONST Fixrec.succeed) (_variable x)"
   247 
   247 
   248 
   248 
   249 subsection {* Pattern combinators for data constructors *}
   249 subsection {* Pattern combinators for data constructors *}
   250 
   250 
   251 types ('a, 'b) pat = "'a \<rightarrow> 'b maybe"
   251 types ('a, 'b) pat = "'a \<rightarrow> 'b match"
   252 
   252 
   253 definition
   253 definition
   254   cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
   254   cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
   255   "cpair_pat p1 p2 = (\<Lambda>(x, y).
   255   "cpair_pat p1 p2 = (\<Lambda>(x, y).
   256     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))"
   256     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))"
   257 
   257 
   258 definition
   258 definition
   259   spair_pat ::
   259   spair_pat ::
   260   "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
   260   "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
   261   "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>(x, y))"
   261   "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>(x, y))"
   272   up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where
   272   up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where
   273   "up_pat p = fup\<cdot>p"
   273   "up_pat p = fup\<cdot>p"
   274 
   274 
   275 definition
   275 definition
   276   TT_pat :: "(tr, unit) pat" where
   276   TT_pat :: "(tr, unit) pat" where
   277   "TT_pat = (\<Lambda> b. If b then return\<cdot>() else fail fi)"
   277   "TT_pat = (\<Lambda> b. If b then succeed\<cdot>() else fail fi)"
   278 
   278 
   279 definition
   279 definition
   280   FF_pat :: "(tr, unit) pat" where
   280   FF_pat :: "(tr, unit) pat" where
   281   "FF_pat = (\<Lambda> b. If b then fail else return\<cdot>() fi)"
   281   "FF_pat = (\<Lambda> b. If b then fail else succeed\<cdot>() fi)"
   282 
   282 
   283 definition
   283 definition
   284   ONE_pat :: "(one, unit) pat" where
   284   ONE_pat :: "(one, unit) pat" where
   285   "ONE_pat = (\<Lambda> ONE. return\<cdot>())"
   285   "ONE_pat = (\<Lambda> ONE. succeed\<cdot>())"
   286 
   286 
   287 text {* Parse translations (patterns) *}
   287 text {* Parse translations (patterns) *}
   288 translations
   288 translations
   289   "_pat (XCONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
   289   "_pat (XCONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
   290   "_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
   290   "_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
   329   "CONST ONE" <= "_match (CONST ONE_pat) _noargs"
   329   "CONST ONE" <= "_match (CONST ONE_pat) _noargs"
   330 
   330 
   331 lemma cpair_pat1:
   331 lemma cpair_pat1:
   332   "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = \<bottom>"
   332   "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = \<bottom>"
   333 apply (simp add: branch_def cpair_pat_def)
   333 apply (simp add: branch_def cpair_pat_def)
   334 apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
   334 apply (cases "p\<cdot>x", simp_all)
   335 done
   335 done
   336 
   336 
   337 lemma cpair_pat2:
   337 lemma cpair_pat2:
   338   "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = fail"
   338   "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = fail"
   339 apply (simp add: branch_def cpair_pat_def)
   339 apply (simp add: branch_def cpair_pat_def)
   340 apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
   340 apply (cases "p\<cdot>x", simp_all)
   341 done
   341 done
   342 
   342 
   343 lemma cpair_pat3:
   343 lemma cpair_pat3:
   344   "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow>
   344   "branch p\<cdot>r\<cdot>x = succeed\<cdot>s \<Longrightarrow>
   345    branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = branch q\<cdot>s\<cdot>y"
   345    branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = branch q\<cdot>s\<cdot>y"
   346 apply (simp add: branch_def cpair_pat_def)
   346 apply (simp add: branch_def cpair_pat_def)
   347 apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
   347 apply (cases "p\<cdot>x", simp_all)
   348 apply (rule_tac p="q\<cdot>y" in maybeE, simp_all)
   348 apply (cases "q\<cdot>y", simp_all)
   349 done
   349 done
   350 
   350 
   351 lemmas cpair_pat [simp] =
   351 lemmas cpair_pat [simp] =
   352   cpair_pat1 cpair_pat2 cpair_pat3
   352   cpair_pat1 cpair_pat2 cpair_pat3
   353 
   353 
   375   "branch (up_pat p)\<cdot>r\<cdot>(up\<cdot>x) = branch p\<cdot>r\<cdot>x"
   375   "branch (up_pat p)\<cdot>r\<cdot>(up\<cdot>x) = branch p\<cdot>r\<cdot>x"
   376 by (simp_all add: branch_def up_pat_def)
   376 by (simp_all add: branch_def up_pat_def)
   377 
   377 
   378 lemma TT_pat [simp]:
   378 lemma TT_pat [simp]:
   379   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   379   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   380   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = return\<cdot>r"
   380   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = succeed\<cdot>r"
   381   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail"
   381   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail"
   382 by (simp_all add: branch_def TT_pat_def)
   382 by (simp_all add: branch_def TT_pat_def)
   383 
   383 
   384 lemma FF_pat [simp]:
   384 lemma FF_pat [simp]:
   385   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   385   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   386   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail"
   386   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail"
   387   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = return\<cdot>r"
   387   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = succeed\<cdot>r"
   388 by (simp_all add: branch_def FF_pat_def)
   388 by (simp_all add: branch_def FF_pat_def)
   389 
   389 
   390 lemma ONE_pat [simp]:
   390 lemma ONE_pat [simp]:
   391   "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   391   "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
   392   "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = return\<cdot>r"
   392   "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = succeed\<cdot>r"
   393 by (simp_all add: branch_def ONE_pat_def)
   393 by (simp_all add: branch_def ONE_pat_def)
   394 
   394 
   395 
   395 
   396 subsection {* Wildcards, as-patterns, and lazy patterns *}
   396 subsection {* Wildcards, as-patterns, and lazy patterns *}
   397 
   397 
   398 definition
   398 definition
   399   wild_pat :: "'a \<rightarrow> unit maybe" where
   399   wild_pat :: "'a \<rightarrow> unit match" where
   400   "wild_pat = (\<Lambda> x. return\<cdot>())"
   400   "wild_pat = (\<Lambda> x. succeed\<cdot>())"
   401 
   401 
   402 definition
   402 definition
   403   as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
   403   as_pat :: "('a \<rightarrow> 'b match) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) match" where
   404   "as_pat p = (\<Lambda> x. maybe_when\<cdot>fail\<cdot>(\<Lambda> a. return\<cdot>(x, a))\<cdot>(p\<cdot>x))"
   404   "as_pat p = (\<Lambda> x. match_case\<cdot>fail\<cdot>(\<Lambda> a. succeed\<cdot>(x, a))\<cdot>(p\<cdot>x))"
   405 
   405 
   406 definition
   406 definition
   407   lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
   407   lazy_pat :: "('a \<rightarrow> 'b::pcpo match) \<Rightarrow> ('a \<rightarrow> 'b match)" where
   408   "lazy_pat p = (\<Lambda> x. return\<cdot>(cases\<cdot>(p\<cdot>x)))"
   408   "lazy_pat p = (\<Lambda> x. succeed\<cdot>(cases\<cdot>(p\<cdot>x)))"
   409 
   409 
   410 text {* Parse translations (patterns) *}
   410 text {* Parse translations (patterns) *}
   411 translations
   411 translations
   412   "_pat _" => "CONST wild_pat"
   412   "_pat _" => "CONST wild_pat"
   413 
   413 
   417 
   417 
   418 text {* Print translations *}
   418 text {* Print translations *}
   419 translations
   419 translations
   420   "_" <= "_match (CONST wild_pat) _noargs"
   420   "_" <= "_match (CONST wild_pat) _noargs"
   421 
   421 
   422 lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = return\<cdot>r"
   422 lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = succeed\<cdot>r"
   423 by (simp add: branch_def wild_pat_def)
   423 by (simp add: branch_def wild_pat_def)
   424 
   424 
   425 lemma as_pat [simp]:
   425 lemma as_pat [simp]:
   426   "branch (as_pat p)\<cdot>(csplit\<cdot>r)\<cdot>x = branch p\<cdot>(r\<cdot>x)\<cdot>x"
   426   "branch (as_pat p)\<cdot>(csplit\<cdot>r)\<cdot>x = branch p\<cdot>(r\<cdot>x)\<cdot>x"
   427 apply (simp add: branch_def as_pat_def)
   427 apply (simp add: branch_def as_pat_def)
   428 apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
   428 apply (cases "p\<cdot>x", simp_all)
   429 done
   429 done
   430 
   430 
   431 lemma lazy_pat [simp]:
   431 lemma lazy_pat [simp]:
   432   "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)"
   432   "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>\<bottom>)"
   433   "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)"
   433   "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>\<bottom>)"
   434   "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>s"
   434   "branch p\<cdot>r\<cdot>x = succeed\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>s"
   435 apply (simp_all add: branch_def lazy_pat_def)
   435 apply (simp_all add: branch_def lazy_pat_def)
   436 apply (rule_tac [!] p="p\<cdot>x" in maybeE, simp_all)
   436 apply (cases "p\<cdot>x", simp_all)+
   437 done
   437 done
   438 
   438 
   439 
   439 
   440 subsection {* Match functions for built-in types *}
   440 subsection {* Match functions for built-in types *}
   441 
   441 
   442 default_sort pcpo
   442 default_sort pcpo
   443 
   443 
   444 definition
   444 definition
   445   match_UU :: "'a \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
   445   match_UU :: "'a \<rightarrow> 'c match \<rightarrow> 'c match"
   446 where
   446 where
   447   "match_UU = strictify\<cdot>(\<Lambda> x k. fail)"
   447   "match_UU = strictify\<cdot>(\<Lambda> x k. fail)"
   448 
   448 
   449 definition
   449 definition
   450   match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   450   match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match"
   451 where
   451 where
   452   "match_cpair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)"
   452   "match_cpair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)"
   453 
   453 
   454 definition
   454 definition
   455   match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   455   match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match"
   456 where
   456 where
   457   "match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)"
   457   "match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)"
   458 
   458 
   459 definition
   459 definition
   460   match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   460   match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match"
   461 where
   461 where
   462   "match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)"
   462   "match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)"
   463 
   463 
   464 definition
   464 definition
   465   match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   465   match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c match) \<rightarrow> 'c match"
   466 where
   466 where
   467   "match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)"
   467   "match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)"
   468 
   468 
   469 definition
   469 definition
   470   match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
   470   match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match"
   471 where
   471 where
   472   "match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)"
   472   "match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)"
   473 
   473 
   474 definition
   474 definition
   475   match_ONE :: "one \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
   475   match_ONE :: "one \<rightarrow> 'c match \<rightarrow> 'c match"
   476 where
   476 where
   477   "match_ONE = (\<Lambda> ONE k. k)"
   477   "match_ONE = (\<Lambda> ONE k. k)"
   478 
   478 
   479 definition
   479 definition
   480   match_TT :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
   480   match_TT :: "tr \<rightarrow> 'c match \<rightarrow> 'c match"
   481 where
   481 where
   482   "match_TT = (\<Lambda> x k. If x then k else fail fi)"
   482   "match_TT = (\<Lambda> x k. If x then k else fail fi)"
   483  
   483  
   484 definition
   484 definition
   485   match_FF :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
   485   match_FF :: "tr \<rightarrow> 'c match \<rightarrow> 'c match"
   486 where
   486 where
   487   "match_FF = (\<Lambda> x k. If x then fail else k fi)"
   487   "match_FF = (\<Lambda> x k. If x then fail else k fi)"
   488 
   488 
   489 lemma match_UU_simps [simp]:
   489 lemma match_UU_simps [simp]:
   490   "match_UU\<cdot>\<bottom>\<cdot>k = \<bottom>"
   490   "match_UU\<cdot>\<bottom>\<cdot>k = \<bottom>"
   582       (@{const_name TT}, @{const_name match_TT}),
   582       (@{const_name TT}, @{const_name match_TT}),
   583       (@{const_name FF}, @{const_name match_FF}),
   583       (@{const_name FF}, @{const_name match_FF}),
   584       (@{const_name UU}, @{const_name match_UU}) ]
   584       (@{const_name UU}, @{const_name match_UU}) ]
   585 *}
   585 *}
   586 
   586 
   587 hide_const (open) return fail run cases
   587 hide_const (open) succeed fail run cases
   588 
   588 
   589 end
   589 end