| author | wenzelm | 
| Sat, 17 May 2008 21:46:24 +0200 | |
| changeset 26933 | 7ca61b1ad872 | 
| parent 26046 | 1624b3304bb9 | 
| child 28891 | f199def7a6a5 | 
| permissions | -rw-r--r-- | 
| 16221 | 1 | (* Title: HOLCF/Fixrec.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Amber Telfer and Brian Huffman | |
| 4 | *) | |
| 5 | ||
| 6 | header "Package for defining recursive functions in HOLCF" | |
| 7 | ||
| 8 | theory Fixrec | |
| 16551 | 9 | imports Sprod Ssum Up One Tr Fix | 
| 23152 | 10 | uses ("Tools/fixrec_package.ML")
 | 
| 16221 | 11 | begin | 
| 12 | ||
| 13 | subsection {* Maybe monad type *}
 | |
| 14 | ||
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 15 | defaultsort cpo | 
| 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 16 | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 17 | pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 18 | by simp | 
| 16221 | 19 | |
| 20 | constdefs | |
| 21 | fail :: "'a maybe" | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 22 | "fail \<equiv> Abs_maybe (sinl\<cdot>ONE)" | 
| 16221 | 23 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 24 | constdefs | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 25 | return :: "'a \<rightarrow> 'a maybe" where | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 26 | "return \<equiv> \<Lambda> x. Abs_maybe (sinr\<cdot>(up\<cdot>x))" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 27 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 28 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 29 |   maybe_when :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a maybe \<rightarrow> 'b::pcpo" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 30 | "maybe_when = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_maybe m))" | 
| 16221 | 31 | |
| 32 | lemma maybeE: | |
| 33 | "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = return\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | |
| 34 | apply (unfold fail_def return_def) | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 35 | apply (cases p, rename_tac r) | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 36 | apply (rule_tac p=r in ssumE, simp add: Abs_maybe_strict) | 
| 16221 | 37 | apply (rule_tac p=x in oneE, simp, simp) | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 38 | apply (rule_tac p=y in upE, simp, simp add: cont_Abs_maybe) | 
| 16221 | 39 | done | 
| 40 | ||
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 41 | lemma return_defined [simp]: "return\<cdot>x \<noteq> \<bottom>" | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 42 | by (simp add: return_def cont_Abs_maybe Abs_maybe_defined) | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 43 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 44 | lemma fail_defined [simp]: "fail \<noteq> \<bottom>" | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 45 | by (simp add: fail_def Abs_maybe_defined) | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 46 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 47 | lemma return_eq [simp]: "(return\<cdot>x = return\<cdot>y) = (x = y)" | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 48 | by (simp add: return_def cont_Abs_maybe Abs_maybe_inject) | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 49 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 50 | lemma return_neq_fail [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 51 | "return\<cdot>x \<noteq> fail" "fail \<noteq> return\<cdot>x" | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 52 | by (simp_all add: return_def fail_def cont_Abs_maybe Abs_maybe_inject) | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 53 | |
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 54 | lemma maybe_when_rews [simp]: | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 55 | "maybe_when\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 56 | "maybe_when\<cdot>f\<cdot>r\<cdot>fail = f" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 57 | "maybe_when\<cdot>f\<cdot>r\<cdot>(return\<cdot>x) = r\<cdot>x" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 58 | by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 59 | cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict) | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 60 | |
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 61 | translations | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 62 | "case m of fail \<Rightarrow> t1 | return\<cdot>x \<Rightarrow> t2" == "CONST maybe_when\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 63 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 64 | |
| 18097 | 65 | subsubsection {* Monadic bind operator *}
 | 
| 16221 | 66 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 67 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 68 |   bind :: "'a maybe \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> 'b maybe" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 69 | "bind = (\<Lambda> m f. case m of fail \<Rightarrow> fail | return\<cdot>x \<Rightarrow> f\<cdot>x)" | 
| 16221 | 70 | |
| 71 | text {* monad laws *}
 | |
| 72 | ||
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 73 | lemma bind_strict [simp]: "bind\<cdot>\<bottom>\<cdot>f = \<bottom>" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 74 | by (simp add: bind_def) | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 75 | |
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 76 | lemma bind_fail [simp]: "bind\<cdot>fail\<cdot>f = fail" | 
| 16221 | 77 | by (simp add: bind_def) | 
| 78 | ||
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 79 | lemma left_unit [simp]: "bind\<cdot>(return\<cdot>a)\<cdot>k = k\<cdot>a" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 80 | by (simp add: bind_def) | 
| 16221 | 81 | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 82 | lemma right_unit [simp]: "bind\<cdot>m\<cdot>return = m" | 
| 16221 | 83 | by (rule_tac p=m in maybeE, simp_all) | 
| 84 | ||
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 85 | lemma bind_assoc: | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 86 | "bind\<cdot>(bind\<cdot>m\<cdot>k)\<cdot>h = bind\<cdot>m\<cdot>(\<Lambda> a. bind\<cdot>(k\<cdot>a)\<cdot>h)" | 
| 16221 | 87 | by (rule_tac p=m in maybeE, simp_all) | 
| 88 | ||
| 18097 | 89 | subsubsection {* Run operator *}
 | 
| 16221 | 90 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 91 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 92 | run:: "'a maybe \<rightarrow> 'a::pcpo" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 93 | "run = maybe_when\<cdot>\<bottom>\<cdot>ID" | 
| 16221 | 94 | |
| 95 | text {* rewrite rules for run *}
 | |
| 96 | ||
| 97 | lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>" | |
| 98 | by (simp add: run_def) | |
| 99 | ||
| 100 | lemma run_fail [simp]: "run\<cdot>fail = \<bottom>" | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 101 | by (simp add: run_def) | 
| 16221 | 102 | |
| 103 | lemma run_return [simp]: "run\<cdot>(return\<cdot>x) = x" | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 104 | by (simp add: run_def) | 
| 16221 | 105 | |
| 18097 | 106 | subsubsection {* Monad plus operator *}
 | 
| 16221 | 107 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 108 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 109 | mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 110 | "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | return\<cdot>x \<Rightarrow> m1)" | 
| 18097 | 111 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 112 | abbreviation | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 113 | mplus_syn :: "['a maybe, 'a maybe] \<Rightarrow> 'a maybe" (infixr "+++" 65) where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 114 | "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2" | 
| 16221 | 115 | |
| 116 | text {* rewrite rules for mplus *}
 | |
| 117 | ||
| 118 | lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>" | |
| 119 | by (simp add: mplus_def) | |
| 120 | ||
| 121 | lemma mplus_fail [simp]: "fail +++ m = m" | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 122 | by (simp add: mplus_def) | 
| 16221 | 123 | |
| 124 | lemma mplus_return [simp]: "return\<cdot>x +++ m = return\<cdot>x" | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 125 | by (simp add: mplus_def) | 
| 16221 | 126 | |
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 127 | lemma mplus_fail2 [simp]: "m +++ fail = m" | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 128 | by (rule_tac p=m in maybeE, simp_all) | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 129 | |
| 16221 | 130 | lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)" | 
| 131 | by (rule_tac p=x in maybeE, simp_all) | |
| 132 | ||
| 18097 | 133 | subsubsection {* Fatbar combinator *}
 | 
| 134 | ||
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 135 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 136 |   fatbar :: "('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe)" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 137 | "fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)" | 
| 18097 | 138 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 139 | abbreviation | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 140 | fatbar_syn :: "['a \<rightarrow> 'b maybe, 'a \<rightarrow> 'b maybe] \<Rightarrow> 'a \<rightarrow> 'b maybe" (infixr "\<parallel>" 60) where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 141 | "m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2" | 
| 18097 | 142 | |
| 143 | lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>" | |
| 144 | by (simp add: fatbar_def) | |
| 145 | ||
| 146 | lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x" | |
| 147 | by (simp add: fatbar_def) | |
| 148 | ||
| 149 | lemma fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = return\<cdot>y" | |
| 150 | by (simp add: fatbar_def) | |
| 151 | ||
| 152 | lemmas fatbar_simps = fatbar1 fatbar2 fatbar3 | |
| 153 | ||
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 154 | lemma run_fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = \<bottom>" | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 155 | by (simp add: fatbar_def) | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 156 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 157 | lemma run_fatbar2: "m\<cdot>x = fail \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = run\<cdot>(ms\<cdot>x)" | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 158 | by (simp add: fatbar_def) | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 159 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 160 | lemma run_fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y" | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 161 | by (simp add: fatbar_def) | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 162 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 163 | lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3 | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 164 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 165 | subsection {* Case branch combinator *}
 | 
| 18097 | 166 | |
| 167 | constdefs | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 168 |   branch :: "('a \<rightarrow> 'b maybe) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c maybe)"
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 169 | "branch p \<equiv> \<Lambda> r x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> y. return\<cdot>(r\<cdot>y))" | 
| 18097 | 170 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 171 | lemma branch_rews: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 172 | "p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 173 | "p\<cdot>x = fail \<Longrightarrow> branch p\<cdot>r\<cdot>x = fail" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 174 | "p\<cdot>x = return\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>y)" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 175 | by (simp_all add: branch_def) | 
| 18097 | 176 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 177 | lemma branch_return [simp]: "branch return\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>x)" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 178 | by (simp add: branch_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 179 | |
| 18097 | 180 | |
| 181 | subsection {* Case syntax *}
 | |
| 182 | ||
| 183 | nonterminals | |
| 184 | Case_syn Cases_syn | |
| 185 | ||
| 186 | syntax | |
| 187 |   "_Case_syntax":: "['a, Cases_syn] => 'b"               ("(Case _ of/ _)" 10)
 | |
| 188 |   "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ =>/ _)" 10)
 | |
| 189 |   ""            :: "Case_syn => Cases_syn"               ("_")
 | |
| 190 |   "_Case2"      :: "[Case_syn, Cases_syn] => Cases_syn"  ("_/ | _")
 | |
| 191 | ||
| 192 | syntax (xsymbols) | |
| 193 |   "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
 | |
| 194 | ||
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 195 | translations | 
| 25158 
271e499f2d03
translations: use XCONST for input patterns (keeps original spelling of const);
 wenzelm parents: 
25131diff
changeset | 196 | "_Case_syntax x ms" == "CONST Fixrec.run\<cdot>(ms\<cdot>x)" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 197 | "_Case2 m ms" == "m \<parallel> ms" | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 198 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 199 | text {* Parsing Case expressions *}
 | 
| 18097 | 200 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 201 | syntax | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 202 | "_pat" :: "'a" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 203 | "_var" :: "'a" | 
| 26046 | 204 | "_noargs" :: "'a" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 205 | |
| 18097 | 206 | translations | 
| 26046 | 207 | "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_var p r)" | 
| 208 | "_var (_args x y) r" => "CONST csplit\<cdot>(_var x (_var y r))" | |
| 209 | "_var _noargs r" => "CONST unit_when\<cdot>r" | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 210 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 211 | parse_translation {*
 | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 212 | (* rewrites (_pat x) => (return) *) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 213 | (* rewrites (_var x t) => (Abs_CFun (%x. t)) *) | 
| 19104 | 214 |   [("_pat", K (Syntax.const "Fixrec.return")),
 | 
| 25158 
271e499f2d03
translations: use XCONST for input patterns (keeps original spelling of const);
 wenzelm parents: 
25131diff
changeset | 215 |    mk_binder_tr ("_var", "Abs_CFun")];
 | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 216 | *} | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 217 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 218 | text {* Printing Case expressions *}
 | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 219 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 220 | syntax | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 221 | "_match" :: "'a" | 
| 18097 | 222 | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 223 | print_translation {*
 | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 224 | let | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 225 |     fun dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax unit_when},_) $ t) =
 | 
| 26046 | 226 | (Syntax.const "_noargs", t) | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 227 |     |   dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax csplit},_) $ t) =
 | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 228 | let | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 229 | val (v1, t1) = dest_LAM t; | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 230 | val (v2, t2) = dest_LAM t1; | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 231 | in (Syntax.const "_args" $ v1 $ v2, t2) end | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 232 |     |   dest_LAM (Const (@{const_syntax Abs_CFun},_) $ t) =
 | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 233 | let | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 234 | val abs = case t of Abs abs => abs | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 235 |                 | _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0);
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 236 | val (x, t') = atomic_abs_tr' abs; | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 237 | in (Syntax.const "_var" $ x, t') end | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 238 | | dest_LAM _ = raise Match; (* too few vars: abort translation *) | 
| 18097 | 239 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 240 |     fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] =
 | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 241 | let val (v, t) = dest_LAM r; | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 242 | in Syntax.const "_Case1" $ (Syntax.const "_match" $ p $ v) $ t end; | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 243 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 244 |   in [(@{const_syntax Rep_CFun}, Case1_tr')] end;
 | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 245 | *} | 
| 18097 | 246 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 247 | translations | 
| 19104 | 248 | "x" <= "_match Fixrec.return (_var x)" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 249 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 250 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 251 | subsection {* Pattern combinators for data constructors *}
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 252 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 253 | types ('a, 'b) pat = "'a \<rightarrow> 'b maybe"
 | 
| 18097 | 254 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 255 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 256 |   cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 257 | "cpair_pat p1 p2 = (\<Lambda>\<langle>x, y\<rangle>. | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 258 | bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>\<langle>a, b\<rangle>)))" | 
| 18097 | 259 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 260 | definition | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 261 | spair_pat :: | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 262 |   "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 263 | "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>\<langle>x, y\<rangle>)" | 
| 18097 | 264 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 265 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 266 |   sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 267 | "sinl_pat p = sscase\<cdot>p\<cdot>(\<Lambda> x. fail)" | 
| 18097 | 268 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 269 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 270 |   sinr_pat :: "('b, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 271 | "sinr_pat p = sscase\<cdot>(\<Lambda> x. fail)\<cdot>p" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 272 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 273 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 274 |   up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 275 | "up_pat p = fup\<cdot>p" | 
| 18097 | 276 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 277 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 278 | TT_pat :: "(tr, unit) pat" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 279 | "TT_pat = (\<Lambda> b. If b then return\<cdot>() else fail fi)" | 
| 18097 | 280 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 281 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 282 | FF_pat :: "(tr, unit) pat" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 283 | "FF_pat = (\<Lambda> b. If b then fail else return\<cdot>() fi)" | 
| 18097 | 284 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 285 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 286 | ONE_pat :: "(one, unit) pat" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 287 | "ONE_pat = (\<Lambda> ONE. return\<cdot>())" | 
| 18097 | 288 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 289 | text {* Parse translations (patterns) *}
 | 
| 18097 | 290 | translations | 
| 26046 | 291 | "_pat (XCONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)" | 
| 292 | "_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)" | |
| 293 | "_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)" | |
| 294 | "_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)" | |
| 295 | "_pat (XCONST up\<cdot>x)" => "CONST up_pat (_pat x)" | |
| 296 | "_pat (XCONST TT)" => "CONST TT_pat" | |
| 297 | "_pat (XCONST FF)" => "CONST FF_pat" | |
| 298 | "_pat (XCONST ONE)" => "CONST ONE_pat" | |
| 299 | ||
| 300 | text {* CONST version is also needed for constructors with special syntax *}
 | |
| 301 | translations | |
| 302 | "_pat (CONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)" | |
| 303 | "_pat (CONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)" | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 304 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 305 | text {* Parse translations (variables) *}
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 306 | translations | 
| 25158 
271e499f2d03
translations: use XCONST for input patterns (keeps original spelling of const);
 wenzelm parents: 
25131diff
changeset | 307 | "_var (XCONST cpair\<cdot>x\<cdot>y) r" => "_var (_args x y) r" | 
| 
271e499f2d03
translations: use XCONST for input patterns (keeps original spelling of const);
 wenzelm parents: 
25131diff
changeset | 308 | "_var (XCONST spair\<cdot>x\<cdot>y) r" => "_var (_args x y) r" | 
| 
271e499f2d03
translations: use XCONST for input patterns (keeps original spelling of const);
 wenzelm parents: 
25131diff
changeset | 309 | "_var (XCONST sinl\<cdot>x) r" => "_var x r" | 
| 
271e499f2d03
translations: use XCONST for input patterns (keeps original spelling of const);
 wenzelm parents: 
25131diff
changeset | 310 | "_var (XCONST sinr\<cdot>x) r" => "_var x r" | 
| 
271e499f2d03
translations: use XCONST for input patterns (keeps original spelling of const);
 wenzelm parents: 
25131diff
changeset | 311 | "_var (XCONST up\<cdot>x) r" => "_var x r" | 
| 26046 | 312 | "_var (XCONST TT) r" => "_var _noargs r" | 
| 313 | "_var (XCONST FF) r" => "_var _noargs r" | |
| 314 | "_var (XCONST ONE) r" => "_var _noargs r" | |
| 315 | ||
| 316 | translations | |
| 317 | "_var (CONST cpair\<cdot>x\<cdot>y) r" => "_var (_args x y) r" | |
| 318 | "_var (CONST spair\<cdot>x\<cdot>y) r" => "_var (_args x y) r" | |
| 18097 | 319 | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 320 | text {* Print translations *}
 | 
| 18097 | 321 | translations | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 322 | "CONST cpair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 323 | <= "_match (CONST cpair_pat p1 p2) (_args v1 v2)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 324 | "CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 325 | <= "_match (CONST spair_pat p1 p2) (_args v1 v2)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 326 | "CONST sinl\<cdot>(_match p1 v1)" <= "_match (CONST sinl_pat p1) v1" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 327 | "CONST sinr\<cdot>(_match p1 v1)" <= "_match (CONST sinr_pat p1) v1" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 328 | "CONST up\<cdot>(_match p1 v1)" <= "_match (CONST up_pat p1) v1" | 
| 26046 | 329 | "CONST TT" <= "_match (CONST TT_pat) _noargs" | 
| 330 | "CONST FF" <= "_match (CONST FF_pat) _noargs" | |
| 331 | "CONST ONE" <= "_match (CONST ONE_pat) _noargs" | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 332 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 333 | lemma cpair_pat1: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 334 | "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 335 | apply (simp add: branch_def cpair_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 336 | apply (rule_tac p="p\<cdot>x" in maybeE, simp_all) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 337 | done | 
| 18097 | 338 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 339 | lemma cpair_pat2: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 340 | "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = fail" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 341 | apply (simp add: branch_def cpair_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 342 | apply (rule_tac p="p\<cdot>x" in maybeE, simp_all) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 343 | done | 
| 18097 | 344 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 345 | lemma cpair_pat3: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 346 | "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow> | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 347 | branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = branch q\<cdot>s\<cdot>y" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 348 | apply (simp add: branch_def cpair_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 349 | apply (rule_tac p="p\<cdot>x" in maybeE, simp_all) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 350 | apply (rule_tac p="q\<cdot>y" in maybeE, simp_all) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 351 | done | 
| 18097 | 352 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 353 | lemmas cpair_pat [simp] = | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 354 | cpair_pat1 cpair_pat2 cpair_pat3 | 
| 18097 | 355 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 356 | lemma spair_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 357 | "branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 358 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 359 | \<Longrightarrow> branch (spair_pat p1 p2)\<cdot>r\<cdot>(:x, y:) = | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 360 | branch (cpair_pat p1 p2)\<cdot>r\<cdot>\<langle>x, y\<rangle>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 361 | by (simp_all add: branch_def spair_pat_def) | 
| 18097 | 362 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 363 | lemma sinl_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 364 | "branch (sinl_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 365 | "x \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = branch p\<cdot>r\<cdot>x" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 366 | "y \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = fail" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 367 | by (simp_all add: branch_def sinl_pat_def) | 
| 18097 | 368 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 369 | lemma sinr_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 370 | "branch (sinr_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 371 | "x \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = fail" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 372 | "y \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = branch p\<cdot>r\<cdot>y" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 373 | by (simp_all add: branch_def sinr_pat_def) | 
| 18097 | 374 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 375 | lemma up_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 376 | "branch (up_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 377 | "branch (up_pat p)\<cdot>r\<cdot>(up\<cdot>x) = branch p\<cdot>r\<cdot>x" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 378 | by (simp_all add: branch_def up_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 379 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 380 | lemma TT_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 381 | "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 382 | "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = return\<cdot>r" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 383 | "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 384 | by (simp_all add: branch_def TT_pat_def) | 
| 18097 | 385 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 386 | lemma FF_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 387 | "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 388 | "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 389 | "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = return\<cdot>r" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 390 | by (simp_all add: branch_def FF_pat_def) | 
| 18097 | 391 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 392 | lemma ONE_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 393 | "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 394 | "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = return\<cdot>r" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 395 | by (simp_all add: branch_def ONE_pat_def) | 
| 18097 | 396 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 397 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 398 | subsection {* Wildcards, as-patterns, and lazy patterns *}
 | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 399 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 400 | syntax | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 401 | "_as_pat" :: "[idt, 'a] \<Rightarrow> 'a" (infixr "\<as>" 10) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 402 |   "_lazy_pat" :: "'a \<Rightarrow> 'a" ("\<lazy> _" [1000] 1000)
 | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 403 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 404 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 405 | wild_pat :: "'a \<rightarrow> unit maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 406 | "wild_pat = (\<Lambda> x. return\<cdot>())" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 407 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 408 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 409 |   as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 410 | "as_pat p = (\<Lambda> x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. return\<cdot>\<langle>x, a\<rangle>))" | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 411 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 412 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 413 |   lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 414 | "lazy_pat p = (\<Lambda> x. return\<cdot>(run\<cdot>(p\<cdot>x)))" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 415 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 416 | text {* Parse translations (patterns) *}
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 417 | translations | 
| 26046 | 418 | "_pat _" => "CONST wild_pat" | 
| 419 | "_pat (_as_pat x y)" => "CONST as_pat (_pat y)" | |
| 420 | "_pat (_lazy_pat x)" => "CONST lazy_pat (_pat x)" | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 421 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 422 | text {* Parse translations (variables) *}
 | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 423 | translations | 
| 26046 | 424 | "_var _ r" => "_var _noargs r" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 425 | "_var (_as_pat x y) r" => "_var (_args x y) r" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 426 | "_var (_lazy_pat x) r" => "_var x r" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 427 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 428 | text {* Print translations *}
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 429 | translations | 
| 26046 | 430 | "_" <= "_match (CONST wild_pat) _noargs" | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 431 | "_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_var x) v)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 432 | "_lazy_pat (_match p v)" <= "_match (CONST lazy_pat p) v" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 433 | |
| 19327 | 434 | text {* Lazy patterns in lambda abstractions *}
 | 
| 435 | translations | |
| 25158 
271e499f2d03
translations: use XCONST for input patterns (keeps original spelling of const);
 wenzelm parents: 
25131diff
changeset | 436 | "_cabs (_lazy_pat p) r" == "CONST Fixrec.run oo (_Case1 (_lazy_pat p) r)" | 
| 19327 | 437 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 438 | lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = return\<cdot>r" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 439 | by (simp add: branch_def wild_pat_def) | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 440 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 441 | lemma as_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 442 | "branch (as_pat p)\<cdot>(csplit\<cdot>r)\<cdot>x = branch p\<cdot>(r\<cdot>x)\<cdot>x" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 443 | apply (simp add: branch_def as_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 444 | apply (rule_tac p="p\<cdot>x" in maybeE, simp_all) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 445 | done | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 446 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 447 | lemma lazy_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 448 | "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 449 | "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 450 | "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>s" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 451 | apply (simp_all add: branch_def lazy_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 452 | apply (rule_tac [!] p="p\<cdot>x" in maybeE, simp_all) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 453 | done | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 454 | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 455 | |
| 16221 | 456 | subsection {* Match functions for built-in types *}
 | 
| 457 | ||
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 458 | defaultsort pcpo | 
| 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 459 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 460 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 461 | match_UU :: "'a \<rightarrow> unit maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 462 | "match_UU = (\<Lambda> x. fail)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 463 | |
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 464 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 465 |   match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<times> 'b) maybe" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 466 | "match_cpair = csplit\<cdot>(\<Lambda> x y. return\<cdot><x,y>)" | 
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 467 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 468 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 469 |   match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<times> 'b) maybe" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 470 | "match_spair = ssplit\<cdot>(\<Lambda> x y. return\<cdot><x,y>)" | 
| 16221 | 471 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 472 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 473 | match_sinl :: "'a \<oplus> 'b \<rightarrow> 'a maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 474 | "match_sinl = sscase\<cdot>return\<cdot>(\<Lambda> y. fail)" | 
| 16551 | 475 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 476 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 477 | match_sinr :: "'a \<oplus> 'b \<rightarrow> 'b maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 478 | "match_sinr = sscase\<cdot>(\<Lambda> x. fail)\<cdot>return" | 
| 16551 | 479 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 480 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 481 | match_up :: "'a::cpo u \<rightarrow> 'a maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 482 | "match_up = fup\<cdot>return" | 
| 16221 | 483 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 484 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 485 | match_ONE :: "one \<rightarrow> unit maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 486 | "match_ONE = (\<Lambda> ONE. return\<cdot>())" | 
| 18094 | 487 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 488 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 489 | match_TT :: "tr \<rightarrow> unit maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 490 | "match_TT = (\<Lambda> b. If b then return\<cdot>() else fail fi)" | 
| 18094 | 491 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 492 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 493 | match_FF :: "tr \<rightarrow> unit maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 494 | "match_FF = (\<Lambda> b. If b then fail else return\<cdot>() fi)" | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 495 | |
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 496 | lemma match_UU_simps [simp]: | 
| 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 497 | "match_UU\<cdot>x = fail" | 
| 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 498 | by (simp add: match_UU_def) | 
| 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 499 | |
| 16221 | 500 | lemma match_cpair_simps [simp]: | 
| 501 | "match_cpair\<cdot><x,y> = return\<cdot><x,y>" | |
| 502 | by (simp add: match_cpair_def) | |
| 503 | ||
| 16551 | 504 | lemma match_spair_simps [simp]: | 
| 505 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> match_spair\<cdot>(:x,y:) = return\<cdot><x,y>" | |
| 506 | "match_spair\<cdot>\<bottom> = \<bottom>" | |
| 507 | by (simp_all add: match_spair_def) | |
| 508 | ||
| 509 | lemma match_sinl_simps [simp]: | |
| 510 | "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinl\<cdot>x) = return\<cdot>x" | |
| 511 | "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinr\<cdot>x) = fail" | |
| 512 | "match_sinl\<cdot>\<bottom> = \<bottom>" | |
| 513 | by (simp_all add: match_sinl_def) | |
| 514 | ||
| 515 | lemma match_sinr_simps [simp]: | |
| 516 | "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinr\<cdot>x) = return\<cdot>x" | |
| 517 | "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinl\<cdot>x) = fail" | |
| 518 | "match_sinr\<cdot>\<bottom> = \<bottom>" | |
| 519 | by (simp_all add: match_sinr_def) | |
| 520 | ||
| 16221 | 521 | lemma match_up_simps [simp]: | 
| 522 | "match_up\<cdot>(up\<cdot>x) = return\<cdot>x" | |
| 523 | "match_up\<cdot>\<bottom> = \<bottom>" | |
| 524 | by (simp_all add: match_up_def) | |
| 525 | ||
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 526 | lemma match_ONE_simps [simp]: | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 527 | "match_ONE\<cdot>ONE = return\<cdot>()" | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 528 | "match_ONE\<cdot>\<bottom> = \<bottom>" | 
| 18094 | 529 | by (simp_all add: match_ONE_def) | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 530 | |
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 531 | lemma match_TT_simps [simp]: | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 532 | "match_TT\<cdot>TT = return\<cdot>()" | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 533 | "match_TT\<cdot>FF = fail" | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 534 | "match_TT\<cdot>\<bottom> = \<bottom>" | 
| 18094 | 535 | by (simp_all add: match_TT_def) | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 536 | |
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 537 | lemma match_FF_simps [simp]: | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 538 | "match_FF\<cdot>FF = return\<cdot>()" | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 539 | "match_FF\<cdot>TT = fail" | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 540 | "match_FF\<cdot>\<bottom> = \<bottom>" | 
| 18094 | 541 | by (simp_all add: match_FF_def) | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 542 | |
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 543 | subsection {* Mutual recursion *}
 | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 544 | |
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 545 | text {*
 | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 546 | The following rules are used to prove unfolding theorems from | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 547 | fixed-point definitions of mutually recursive functions. | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 548 | *} | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 549 | |
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 550 | lemma cpair_equalI: "\<lbrakk>x \<equiv> cfst\<cdot>p; y \<equiv> csnd\<cdot>p\<rbrakk> \<Longrightarrow> <x,y> \<equiv> p" | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 551 | by (simp add: surjective_pairing_Cprod2) | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 552 | |
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 553 | lemma cpair_eqD1: "<x,y> = <x',y'> \<Longrightarrow> x = x'" | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 554 | by simp | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 555 | |
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 556 | lemma cpair_eqD2: "<x,y> = <x',y'> \<Longrightarrow> y = y'" | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 557 | by simp | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 558 | |
| 16463 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 559 | text {* lemma for proving rewrite rules *}
 | 
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 560 | |
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 561 | lemma ssubst_lhs: "\<lbrakk>t = s; P s = Q\<rbrakk> \<Longrightarrow> P t = Q" | 
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 562 | by simp | 
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 563 | |
| 16221 | 564 | |
| 16758 | 565 | subsection {* Initializing the fixrec package *}
 | 
| 16221 | 566 | |
| 23152 | 567 | use "Tools/fixrec_package.ML" | 
| 16221 | 568 | |
| 19439 | 569 | hide (open) const return bind fail run | 
| 19104 | 570 | |
| 16221 | 571 | end |