| author | wenzelm | 
| Thu, 18 Feb 2010 20:44:22 +0100 | |
| changeset 35199 | 2e37cdae7b9c | 
| parent 35115 | 446c5063e4fd | 
| child 35469 | 6e59de61d501 | 
| permissions | -rw-r--r-- | 
| 16221 | 1 | (* Title: HOLCF/Fixrec.thy | 
| 2 | Author: Amber Telfer and Brian Huffman | |
| 3 | *) | |
| 4 | ||
| 5 | header "Package for defining recursive functions in HOLCF" | |
| 6 | ||
| 7 | theory Fixrec | |
| 16551 | 8 | imports Sprod Ssum Up One Tr Fix | 
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31095diff
changeset | 9 | uses ("Tools/fixrec.ML")
 | 
| 16221 | 10 | begin | 
| 11 | ||
| 12 | subsection {* Maybe monad type *}
 | |
| 13 | ||
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 14 | defaultsort cpo | 
| 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 15 | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 16 | pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set" | 
| 29063 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
28891diff
changeset | 17 | by simp_all | 
| 16221 | 18 | |
| 29141 | 19 | definition | 
| 20 | fail :: "'a maybe" where | |
| 21 | "fail = Abs_maybe (sinl\<cdot>ONE)" | |
| 16221 | 22 | |
| 29141 | 23 | definition | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 24 | return :: "'a \<rightarrow> 'a maybe" where | 
| 29141 | 25 | "return = (\<Lambda> x. Abs_maybe (sinr\<cdot>(up\<cdot>x)))" | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 26 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 27 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 28 |   maybe_when :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a maybe \<rightarrow> 'b::pcpo" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 29 | "maybe_when = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_maybe m))" | 
| 16221 | 30 | |
| 31 | lemma maybeE: | |
| 32 | "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = return\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | |
| 33 | apply (unfold fail_def return_def) | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 34 | apply (cases p, rename_tac r) | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 35 | apply (rule_tac p=r in ssumE, simp add: Abs_maybe_strict) | 
| 16221 | 36 | 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 | 37 | apply (rule_tac p=y in upE, simp, simp add: cont_Abs_maybe) | 
| 16221 | 38 | done | 
| 39 | ||
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 40 | 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 | 41 | 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 | 42 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 43 | lemma fail_defined [simp]: "fail \<noteq> \<bottom>" | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 44 | by (simp add: fail_def Abs_maybe_defined) | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 45 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 46 | 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 | 47 | 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 | 48 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 49 | lemma return_neq_fail [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 50 | "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 | 51 | 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 | 52 | |
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 53 | lemma maybe_when_rews [simp]: | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 54 | "maybe_when\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 55 | "maybe_when\<cdot>f\<cdot>r\<cdot>fail = f" | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 56 | "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 | 57 | by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe | 
| 29530 
9905b660612b
change to simpler, more extensible continuity simproc
 huffman parents: 
29322diff
changeset | 58 | cont2cont_LAM | 
| 19092 
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 | 
| 29141 | 62 | "case m of XCONST fail \<Rightarrow> t1 | XCONST return\<cdot>x \<Rightarrow> t2" | 
| 63 | == "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 | 64 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 65 | |
| 18097 | 66 | subsubsection {* Monadic bind operator *}
 | 
| 16221 | 67 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 68 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 69 |   bind :: "'a maybe \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> 'b maybe" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 70 | "bind = (\<Lambda> m f. case m of fail \<Rightarrow> fail | return\<cdot>x \<Rightarrow> f\<cdot>x)" | 
| 16221 | 71 | |
| 72 | text {* monad laws *}
 | |
| 73 | ||
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 74 | 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 | 75 | by (simp add: bind_def) | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 76 | |
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 77 | lemma bind_fail [simp]: "bind\<cdot>fail\<cdot>f = fail" | 
| 16221 | 78 | by (simp add: bind_def) | 
| 79 | ||
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 80 | 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 | 81 | by (simp add: bind_def) | 
| 16221 | 82 | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 83 | lemma right_unit [simp]: "bind\<cdot>m\<cdot>return = m" | 
| 16221 | 84 | by (rule_tac p=m in maybeE, simp_all) | 
| 85 | ||
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 86 | lemma bind_assoc: | 
| 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 87 | "bind\<cdot>(bind\<cdot>m\<cdot>k)\<cdot>h = bind\<cdot>m\<cdot>(\<Lambda> a. bind\<cdot>(k\<cdot>a)\<cdot>h)" | 
| 16221 | 88 | by (rule_tac p=m in maybeE, simp_all) | 
| 89 | ||
| 18097 | 90 | subsubsection {* Run operator *}
 | 
| 16221 | 91 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 92 | definition | 
| 28891 | 93 | run :: "'a maybe \<rightarrow> 'a::pcpo" where | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 94 | "run = maybe_when\<cdot>\<bottom>\<cdot>ID" | 
| 16221 | 95 | |
| 96 | text {* rewrite rules for run *}
 | |
| 97 | ||
| 98 | lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>" | |
| 99 | by (simp add: run_def) | |
| 100 | ||
| 101 | lemma run_fail [simp]: "run\<cdot>fail = \<bottom>" | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 102 | by (simp add: run_def) | 
| 16221 | 103 | |
| 104 | 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 | 105 | by (simp add: run_def) | 
| 16221 | 106 | |
| 18097 | 107 | subsubsection {* Monad plus operator *}
 | 
| 16221 | 108 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 109 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 110 | mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 111 | "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | return\<cdot>x \<Rightarrow> m1)" | 
| 18097 | 112 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 113 | abbreviation | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 114 | mplus_syn :: "['a maybe, 'a maybe] \<Rightarrow> 'a maybe" (infixr "+++" 65) where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 115 | "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2" | 
| 16221 | 116 | |
| 117 | text {* rewrite rules for mplus *}
 | |
| 118 | ||
| 119 | lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>" | |
| 120 | by (simp add: mplus_def) | |
| 121 | ||
| 122 | lemma mplus_fail [simp]: "fail +++ m = m" | |
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 123 | by (simp add: mplus_def) | 
| 16221 | 124 | |
| 125 | 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 | 126 | by (simp add: mplus_def) | 
| 16221 | 127 | |
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 128 | lemma mplus_fail2 [simp]: "m +++ fail = m" | 
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 129 | 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 | 130 | |
| 16221 | 131 | lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)" | 
| 132 | by (rule_tac p=x in maybeE, simp_all) | |
| 133 | ||
| 18097 | 134 | subsubsection {* Fatbar combinator *}
 | 
| 135 | ||
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 136 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 137 |   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 | 138 | "fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)" | 
| 18097 | 139 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 140 | abbreviation | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 141 | 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 | 142 | "m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2" | 
| 18097 | 143 | |
| 144 | lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>" | |
| 145 | by (simp add: fatbar_def) | |
| 146 | ||
| 147 | lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x" | |
| 148 | by (simp add: fatbar_def) | |
| 149 | ||
| 150 | lemma fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = return\<cdot>y" | |
| 151 | by (simp add: fatbar_def) | |
| 152 | ||
| 153 | lemmas fatbar_simps = fatbar1 fatbar2 fatbar3 | |
| 154 | ||
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 155 | 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 | 156 | by (simp add: fatbar_def) | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 157 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 158 | 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 | 159 | by (simp add: fatbar_def) | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 160 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 161 | 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 | 162 | by (simp add: fatbar_def) | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 163 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 164 | 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 | 165 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 166 | subsection {* Case branch combinator *}
 | 
| 18097 | 167 | |
| 29141 | 168 | definition | 
| 169 |   branch :: "('a \<rightarrow> 'b maybe) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c maybe)" where
 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 170 | "branch p \<equiv> \<Lambda> r x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> y. return\<cdot>(r\<cdot>y))" | 
| 18097 | 171 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 172 | lemma branch_rews: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 173 | "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 | 174 | "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 | 175 | "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 | 176 | by (simp_all add: branch_def) | 
| 18097 | 177 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 178 | 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 | 179 | by (simp add: branch_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 180 | |
| 28891 | 181 | subsubsection {* Cases operator *}
 | 
| 182 | ||
| 183 | definition | |
| 184 | cases :: "'a maybe \<rightarrow> 'a::pcpo" where | |
| 185 | "cases = maybe_when\<cdot>\<bottom>\<cdot>ID" | |
| 186 | ||
| 187 | text {* rewrite rules for cases *}
 | |
| 188 | ||
| 189 | lemma cases_strict [simp]: "cases\<cdot>\<bottom> = \<bottom>" | |
| 190 | by (simp add: cases_def) | |
| 191 | ||
| 192 | lemma cases_fail [simp]: "cases\<cdot>fail = \<bottom>" | |
| 193 | by (simp add: cases_def) | |
| 194 | ||
| 195 | lemma cases_return [simp]: "cases\<cdot>(return\<cdot>x) = x" | |
| 196 | by (simp add: cases_def) | |
| 18097 | 197 | |
| 198 | subsection {* Case syntax *}
 | |
| 199 | ||
| 200 | nonterminals | |
| 201 | Case_syn Cases_syn | |
| 202 | ||
| 203 | syntax | |
| 204 |   "_Case_syntax":: "['a, Cases_syn] => 'b"               ("(Case _ of/ _)" 10)
 | |
| 205 |   "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ =>/ _)" 10)
 | |
| 206 |   ""            :: "Case_syn => Cases_syn"               ("_")
 | |
| 207 |   "_Case2"      :: "[Case_syn, Cases_syn] => Cases_syn"  ("_/ | _")
 | |
| 208 | ||
| 209 | syntax (xsymbols) | |
| 210 |   "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
 | |
| 211 | ||
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 212 | translations | 
| 28891 | 213 | "_Case_syntax x ms" == "CONST Fixrec.cases\<cdot>(ms\<cdot>x)" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 214 | "_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 | 215 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 216 | text {* Parsing Case expressions *}
 | 
| 18097 | 217 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 218 | syntax | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 219 | "_pat" :: "'a" | 
| 29322 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 220 | "_variable" :: "'a" | 
| 26046 | 221 | "_noargs" :: "'a" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 222 | |
| 18097 | 223 | translations | 
| 29322 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 224 | "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 225 | "_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 226 | "_variable _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 | 227 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 228 | parse_translation {*
 | 
| 35115 | 229 | (* rewrite (_pat x) => (return) *) | 
| 230 | (* rewrite (_variable x t) => (Abs_CFun (%x. t)) *) | |
| 231 |  [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.return}),
 | |
| 232 |   mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_CFun})];
 | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 233 | *} | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 234 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 235 | text {* Printing Case expressions *}
 | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 236 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 237 | syntax | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 238 | "_match" :: "'a" | 
| 18097 | 239 | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 240 | print_translation {*
 | 
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 241 | let | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 242 |     fun dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax unit_when},_) $ t) =
 | 
| 35115 | 243 |           (Syntax.const @{syntax_const "_noargs"}, t)
 | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 244 |     |   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 | 245 | let | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 246 | val (v1, t1) = dest_LAM t; | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 247 | val (v2, t2) = dest_LAM t1; | 
| 35115 | 248 |           in (Syntax.const @{syntax_const "_args"} $ v1 $ v2, t2) end
 | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 249 |     |   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 | 250 | let | 
| 35115 | 251 | val abs = | 
| 252 | case t of Abs abs => abs | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 253 |                 | _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0);
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 254 | val (x, t') = atomic_abs_tr' abs; | 
| 35115 | 255 |           in (Syntax.const @{syntax_const "_variable"} $ x, t') end
 | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 256 | | dest_LAM _ = raise Match; (* too few vars: abort translation *) | 
| 18097 | 257 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 258 |     fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] =
 | 
| 35115 | 259 | let val (v, t) = dest_LAM r in | 
| 260 |             Syntax.const @{syntax_const "_Case1"} $
 | |
| 261 |               (Syntax.const @{syntax_const "_match"} $ p $ v) $ t
 | |
| 262 | end; | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 263 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 264 |   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 | 265 | *} | 
| 18097 | 266 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 267 | translations | 
| 29322 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 268 | "x" <= "_match Fixrec.return (_variable x)" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 269 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 270 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 271 | subsection {* Pattern combinators for data constructors *}
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 272 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 273 | types ('a, 'b) pat = "'a \<rightarrow> 'b maybe"
 | 
| 18097 | 274 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 275 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 276 |   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 | 277 | "cpair_pat p1 p2 = (\<Lambda>\<langle>x, y\<rangle>. | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 278 | bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>\<langle>a, b\<rangle>)))" | 
| 18097 | 279 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 280 | definition | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 281 | spair_pat :: | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 282 |   "('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 | 283 | "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>\<langle>x, y\<rangle>)" | 
| 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 |   sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 287 | "sinl_pat p = sscase\<cdot>p\<cdot>(\<Lambda> x. fail)" | 
| 18097 | 288 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 289 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 290 |   sinr_pat :: "('b, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 291 | "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 | 292 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 293 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 294 |   up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 295 | "up_pat p = fup\<cdot>p" | 
| 18097 | 296 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 297 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 298 | TT_pat :: "(tr, unit) pat" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 299 | "TT_pat = (\<Lambda> b. If b then return\<cdot>() else fail fi)" | 
| 18097 | 300 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 301 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 302 | FF_pat :: "(tr, unit) pat" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 303 | "FF_pat = (\<Lambda> b. If b then fail else return\<cdot>() fi)" | 
| 18097 | 304 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 305 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 306 | ONE_pat :: "(one, unit) pat" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 307 | "ONE_pat = (\<Lambda> ONE. return\<cdot>())" | 
| 18097 | 308 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 309 | text {* Parse translations (patterns) *}
 | 
| 18097 | 310 | translations | 
| 26046 | 311 | "_pat (XCONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)" | 
| 312 | "_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)" | |
| 313 | "_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)" | |
| 314 | "_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)" | |
| 315 | "_pat (XCONST up\<cdot>x)" => "CONST up_pat (_pat x)" | |
| 316 | "_pat (XCONST TT)" => "CONST TT_pat" | |
| 317 | "_pat (XCONST FF)" => "CONST FF_pat" | |
| 318 | "_pat (XCONST ONE)" => "CONST ONE_pat" | |
| 319 | ||
| 320 | text {* CONST version is also needed for constructors with special syntax *}
 | |
| 321 | translations | |
| 322 | "_pat (CONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)" | |
| 323 | "_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 | 324 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 325 | text {* Parse translations (variables) *}
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 326 | translations | 
| 29322 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 327 | "_variable (XCONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 328 | "_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 329 | "_variable (XCONST sinl\<cdot>x) r" => "_variable x r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 330 | "_variable (XCONST sinr\<cdot>x) r" => "_variable x r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 331 | "_variable (XCONST up\<cdot>x) r" => "_variable x r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 332 | "_variable (XCONST TT) r" => "_variable _noargs r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 333 | "_variable (XCONST FF) r" => "_variable _noargs r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 334 | "_variable (XCONST ONE) r" => "_variable _noargs r" | 
| 26046 | 335 | |
| 336 | translations | |
| 29322 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 337 | "_variable (CONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 338 | "_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r" | 
| 18097 | 339 | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 340 | text {* Print translations *}
 | 
| 18097 | 341 | translations | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 342 | "CONST cpair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 343 | <= "_match (CONST cpair_pat p1 p2) (_args v1 v2)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 344 | "CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 345 | <= "_match (CONST spair_pat p1 p2) (_args v1 v2)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 346 | "CONST sinl\<cdot>(_match p1 v1)" <= "_match (CONST sinl_pat p1) v1" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 347 | "CONST sinr\<cdot>(_match p1 v1)" <= "_match (CONST sinr_pat p1) v1" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 348 | "CONST up\<cdot>(_match p1 v1)" <= "_match (CONST up_pat p1) v1" | 
| 26046 | 349 | "CONST TT" <= "_match (CONST TT_pat) _noargs" | 
| 350 | "CONST FF" <= "_match (CONST FF_pat) _noargs" | |
| 351 | "CONST ONE" <= "_match (CONST ONE_pat) _noargs" | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 352 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 353 | lemma cpair_pat1: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 354 | "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 | 355 | apply (simp add: branch_def cpair_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 356 | 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 | 357 | done | 
| 18097 | 358 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 359 | lemma cpair_pat2: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 360 | "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 | 361 | apply (simp add: branch_def cpair_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 362 | 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 | 363 | done | 
| 18097 | 364 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 365 | lemma cpair_pat3: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 366 | "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow> | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 367 | 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 | 368 | apply (simp add: branch_def cpair_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 369 | 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 | 370 | 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 | 371 | done | 
| 18097 | 372 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 373 | lemmas cpair_pat [simp] = | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 374 | cpair_pat1 cpair_pat2 cpair_pat3 | 
| 18097 | 375 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 376 | lemma spair_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 377 | "branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 378 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 379 | \<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 | 380 | 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 | 381 | by (simp_all add: branch_def spair_pat_def) | 
| 18097 | 382 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 383 | lemma sinl_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 384 | "branch (sinl_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 385 | "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 | 386 | "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 | 387 | by (simp_all add: branch_def sinl_pat_def) | 
| 18097 | 388 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 389 | lemma sinr_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 390 | "branch (sinr_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 391 | "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 | 392 | "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 | 393 | by (simp_all add: branch_def sinr_pat_def) | 
| 18097 | 394 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 395 | lemma up_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 396 | "branch (up_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 397 | "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 | 398 | by (simp_all add: branch_def up_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 399 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 400 | lemma TT_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 401 | "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 | 402 | "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 | 403 | "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 | 404 | by (simp_all add: branch_def TT_pat_def) | 
| 18097 | 405 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 406 | lemma FF_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 407 | "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 | 408 | "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 | 409 | "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 | 410 | by (simp_all add: branch_def FF_pat_def) | 
| 18097 | 411 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 412 | lemma ONE_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 413 | "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 | 414 | "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 | 415 | by (simp_all add: branch_def ONE_pat_def) | 
| 18097 | 416 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 417 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 418 | 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 | 419 | |
| 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 420 | syntax | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 421 | "_as_pat" :: "[idt, 'a] \<Rightarrow> 'a" (infixr "\<as>" 10) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 422 |   "_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 | 423 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 424 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 425 | wild_pat :: "'a \<rightarrow> unit maybe" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 426 | "wild_pat = (\<Lambda> x. return\<cdot>())" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 427 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 428 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 429 |   as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 430 | "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 | 431 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 432 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 433 |   lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
 | 
| 28891 | 434 | "lazy_pat p = (\<Lambda> x. return\<cdot>(cases\<cdot>(p\<cdot>x)))" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 435 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 436 | text {* Parse translations (patterns) *}
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 437 | translations | 
| 26046 | 438 | "_pat _" => "CONST wild_pat" | 
| 439 | "_pat (_as_pat x y)" => "CONST as_pat (_pat y)" | |
| 440 | "_pat (_lazy_pat x)" => "CONST lazy_pat (_pat x)" | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 441 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 442 | text {* Parse translations (variables) *}
 | 
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 443 | translations | 
| 29322 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 444 | "_variable _ r" => "_variable _noargs r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 445 | "_variable (_as_pat x y) r" => "_variable (_args x y) r" | 
| 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 446 | "_variable (_lazy_pat x) r" => "_variable x r" | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 447 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 448 | text {* Print translations *}
 | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 449 | translations | 
| 26046 | 450 | "_" <= "_match (CONST wild_pat) _noargs" | 
| 29322 
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
 wenzelm parents: 
29141diff
changeset | 451 | "_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_variable x) v)" | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 452 | "_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 | 453 | |
| 19327 | 454 | text {* Lazy patterns in lambda abstractions *}
 | 
| 455 | translations | |
| 28891 | 456 | "_cabs (_lazy_pat p) r" == "CONST Fixrec.cases oo (_Case1 (_lazy_pat p) r)" | 
| 19327 | 457 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 458 | 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 | 459 | 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 | 460 | |
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 461 | lemma as_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 462 | "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 | 463 | apply (simp add: branch_def as_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 464 | 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 | 465 | done | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 466 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 467 | lemma lazy_pat [simp]: | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 468 | "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 | 469 | "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 | 470 | "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 | 471 | apply (simp_all add: branch_def lazy_pat_def) | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 472 | 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 | 473 | done | 
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 474 | |
| 18112 
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
 huffman parents: 
18110diff
changeset | 475 | |
| 16221 | 476 | subsection {* Match functions for built-in types *}
 | 
| 477 | ||
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 478 | defaultsort pcpo | 
| 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 479 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 480 | definition | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 481 | match_UU :: "'a \<rightarrow> 'c maybe \<rightarrow> 'c maybe" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 482 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 483 | "match_UU = strictify\<cdot>(\<Lambda> x k. fail)" | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 484 | |
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 485 | definition | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 486 |   match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
 | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 487 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 488 | "match_cpair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)" | 
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 489 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 490 | definition | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 491 |   match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
 | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 492 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 493 | "match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)" | 
| 16221 | 494 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 495 | definition | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 496 |   match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
 | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 497 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 498 | "match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)" | 
| 16551 | 499 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 500 | definition | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 501 |   match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
 | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 502 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 503 | "match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)" | 
| 16551 | 504 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 505 | definition | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 506 |   match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
 | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 507 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 508 | "match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)" | 
| 16221 | 509 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 510 | definition | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 511 | match_ONE :: "one \<rightarrow> 'c maybe \<rightarrow> 'c maybe" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 512 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 513 | "match_ONE = (\<Lambda> ONE k. k)" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 514 | |
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 515 | definition | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 516 | match_TT :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 517 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 518 | "match_TT = (\<Lambda> x k. If x then k else fail fi)" | 
| 18094 | 519 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 520 | definition | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 521 | match_FF :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 522 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 523 | "match_FF = (\<Lambda> x k. If x then fail else k fi)" | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 524 | |
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 525 | lemma match_UU_simps [simp]: | 
| 31008 
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
 huffman parents: 
30914diff
changeset | 526 | "match_UU\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
 huffman parents: 
30914diff
changeset | 527 | "x \<noteq> \<bottom> \<Longrightarrow> match_UU\<cdot>x\<cdot>k = fail" | 
| 
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
 huffman parents: 
30914diff
changeset | 528 | by (simp_all add: match_UU_def) | 
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 529 | |
| 16221 | 530 | lemma match_cpair_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 531 | "match_cpair\<cdot>\<langle>x, y\<rangle>\<cdot>k = k\<cdot>x\<cdot>y" | 
| 33401 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
31738diff
changeset | 532 | "match_cpair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y" | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
31738diff
changeset | 533 | by (simp_all add: match_cpair_def) | 
| 16221 | 534 | |
| 16551 | 535 | lemma match_spair_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 536 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> match_spair\<cdot>(:x, y:)\<cdot>k = k\<cdot>x\<cdot>y" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 537 | "match_spair\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 16551 | 538 | by (simp_all add: match_spair_def) | 
| 539 | ||
| 540 | lemma match_sinl_simps [simp]: | |
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 541 | "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinl\<cdot>x)\<cdot>k = k\<cdot>x" | 
| 30914 
ceeb5f2eac75
fix too-specific types in lemmas match_{sinl,sinr}_simps
 huffman parents: 
30912diff
changeset | 542 | "y \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinr\<cdot>y)\<cdot>k = fail" | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 543 | "match_sinl\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 16551 | 544 | by (simp_all add: match_sinl_def) | 
| 545 | ||
| 546 | lemma match_sinr_simps [simp]: | |
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 547 | "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinl\<cdot>x)\<cdot>k = fail" | 
| 30914 
ceeb5f2eac75
fix too-specific types in lemmas match_{sinl,sinr}_simps
 huffman parents: 
30912diff
changeset | 548 | "y \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinr\<cdot>y)\<cdot>k = k\<cdot>y" | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 549 | "match_sinr\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 16551 | 550 | by (simp_all add: match_sinr_def) | 
| 551 | ||
| 16221 | 552 | lemma match_up_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 553 | "match_up\<cdot>(up\<cdot>x)\<cdot>k = k\<cdot>x" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 554 | "match_up\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 16221 | 555 | by (simp_all add: match_up_def) | 
| 556 | ||
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 557 | lemma match_ONE_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 558 | "match_ONE\<cdot>ONE\<cdot>k = k" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 559 | "match_ONE\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 18094 | 560 | 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 | 561 | |
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 562 | lemma match_TT_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 563 | "match_TT\<cdot>TT\<cdot>k = k" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 564 | "match_TT\<cdot>FF\<cdot>k = fail" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 565 | "match_TT\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 18094 | 566 | 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 | 567 | |
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 568 | lemma match_FF_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 569 | "match_FF\<cdot>FF\<cdot>k = k" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 570 | "match_FF\<cdot>TT\<cdot>k = fail" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 571 | "match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 18094 | 572 | 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 | 573 | |
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 574 | subsection {* Mutual recursion *}
 | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 575 | |
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 576 | text {*
 | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 577 | The following rules are used to prove unfolding theorems from | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 578 | fixed-point definitions of mutually recursive functions. | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 579 | *} | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 580 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 581 | lemma Pair_equalI: "\<lbrakk>x \<equiv> fst p; y \<equiv> snd p\<rbrakk> \<Longrightarrow> (x, y) \<equiv> p" | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 582 | by simp | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 583 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 584 | lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'" | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 585 | by simp | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 586 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 587 | lemma Pair_eqD2: "(x, y) = (x', y') \<Longrightarrow> y = y'" | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 588 | by simp | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 589 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 590 | lemma def_cont_fix_eq: | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 591 | "\<lbrakk>f \<equiv> fix\<cdot>(Abs_CFun F); cont F\<rbrakk> \<Longrightarrow> f = F f" | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 592 | by (simp, subst fix_eq, simp) | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 593 | |
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 594 | lemma def_cont_fix_ind: | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 595 | "\<lbrakk>f \<equiv> fix\<cdot>(Abs_CFun F); cont F; adm P; P \<bottom>; \<And>x. P x \<Longrightarrow> P (F x)\<rbrakk> \<Longrightarrow> P f" | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 596 | by (simp add: fix_ind) | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 597 | |
| 16463 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 598 | text {* lemma for proving rewrite rules *}
 | 
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 599 | |
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 600 | 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 | 601 | by simp | 
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 602 | |
| 16221 | 603 | |
| 16758 | 604 | subsection {* Initializing the fixrec package *}
 | 
| 16221 | 605 | |
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31095diff
changeset | 606 | use "Tools/fixrec.ML" | 
| 16221 | 607 | |
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31095diff
changeset | 608 | setup {* Fixrec.setup *}
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 609 | |
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 610 | setup {*
 | 
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31095diff
changeset | 611 | Fixrec.add_matchers | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 612 |     [ (@{const_name up}, @{const_name match_up}),
 | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 613 |       (@{const_name sinl}, @{const_name match_sinl}),
 | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 614 |       (@{const_name sinr}, @{const_name match_sinr}),
 | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 615 |       (@{const_name spair}, @{const_name match_spair}),
 | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 616 |       (@{const_name cpair}, @{const_name match_cpair}),
 | 
| 33401 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
31738diff
changeset | 617 |       (@{const_name Pair}, @{const_name match_cpair}),
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 618 |       (@{const_name ONE}, @{const_name match_ONE}),
 | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 619 |       (@{const_name TT}, @{const_name match_TT}),
 | 
| 31008 
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
 huffman parents: 
30914diff
changeset | 620 |       (@{const_name FF}, @{const_name match_FF}),
 | 
| 
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
 huffman parents: 
30914diff
changeset | 621 |       (@{const_name UU}, @{const_name match_UU}) ]
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 622 | *} | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 623 | |
| 28891 | 624 | hide (open) const return bind fail run cases | 
| 19104 | 625 | |
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 626 | lemmas [fixrec_simp] = | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 627 | run_strict run_fail run_return | 
| 33429 | 628 | mplus_strict mplus_fail mplus_return | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 629 | spair_strict_iff | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 630 | sinl_defined_iff | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 631 | sinr_defined_iff | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 632 | up_defined | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 633 | ONE_defined | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 634 | dist_eq_tr(1,2) | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 635 | match_UU_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 636 | match_cpair_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 637 | match_spair_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 638 | match_sinl_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 639 | match_sinr_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 640 | match_up_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 641 | match_ONE_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 642 | match_TT_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 643 | match_FF_simps | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 644 | |
| 16221 | 645 | end |