author | blanchet |
Wed, 20 Jan 2010 10:38:19 +0100 | |
changeset 34937 | fb90752a9271 |
parent 33429 | 42d7b6b4992b |
child 35115 | 446c5063e4fd |
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:
31095
diff
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:
16758
diff
changeset
|
14 |
defaultsort cpo |
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
15 |
|
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
16 |
pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set" |
29063
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents:
28891
diff
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:
23152
diff
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:
18293
diff
changeset
|
26 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
27 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
28 |
maybe_when :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a maybe \<rightarrow> 'b::pcpo" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
18293
diff
changeset
|
34 |
apply (cases p, rename_tac r) |
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
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:
18293
diff
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:
18112
diff
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:
18293
diff
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:
18112
diff
changeset
|
42 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
43 |
lemma fail_defined [simp]: "fail \<noteq> \<bottom>" |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
44 |
by (simp add: fail_def Abs_maybe_defined) |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
45 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
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:
18293
diff
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:
18112
diff
changeset
|
48 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
49 |
lemma return_neq_fail [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
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:
18293
diff
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:
18293
diff
changeset
|
52 |
|
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
53 |
lemma maybe_when_rews [simp]: |
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
54 |
"maybe_when\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>" |
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
55 |
"maybe_when\<cdot>f\<cdot>r\<cdot>fail = f" |
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
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:
18293
diff
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:
29322
diff
changeset
|
58 |
cont2cont_LAM |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
59 |
cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict) |
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
60 |
|
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
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:
18112
diff
changeset
|
64 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
65 |
|
18097 | 66 |
subsubsection {* Monadic bind operator *} |
16221 | 67 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
68 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
69 |
bind :: "'a maybe \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> 'b maybe" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
18293
diff
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:
18293
diff
changeset
|
75 |
by (simp add: bind_def) |
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
76 |
|
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
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:
18293
diff
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:
18293
diff
changeset
|
81 |
by (simp add: bind_def) |
16221 | 82 |
|
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
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:
18293
diff
changeset
|
86 |
lemma bind_assoc: |
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
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:
23152
diff
changeset
|
92 |
definition |
28891 | 93 |
run :: "'a maybe \<rightarrow> 'a::pcpo" where |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
18293
diff
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:
18293
diff
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:
23152
diff
changeset
|
109 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
110 |
mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
23152
diff
changeset
|
113 |
abbreviation |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
114 |
mplus_syn :: "['a maybe, 'a maybe] \<Rightarrow> 'a maybe" (infixr "+++" 65) where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
18293
diff
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:
18293
diff
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:
16417
diff
changeset
|
128 |
lemma mplus_fail2 [simp]: "m +++ fail = m" |
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
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:
16417
diff
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:
23152
diff
changeset
|
136 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
23152
diff
changeset
|
138 |
"fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)" |
18097 | 139 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
140 |
abbreviation |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
23152
diff
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:
18110
diff
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:
18110
diff
changeset
|
156 |
by (simp add: fatbar_def) |
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
157 |
|
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
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:
18110
diff
changeset
|
159 |
by (simp add: fatbar_def) |
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
160 |
|
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
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:
18110
diff
changeset
|
162 |
by (simp add: fatbar_def) |
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
163 |
|
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
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:
18110
diff
changeset
|
165 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
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:
18112
diff
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:
18112
diff
changeset
|
172 |
lemma branch_rews: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
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:
18112
diff
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:
18112
diff
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:
18112
diff
changeset
|
176 |
by (simp_all add: branch_def) |
18097 | 177 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
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:
18112
diff
changeset
|
179 |
by (simp add: branch_def) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
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:
18112
diff
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:
18112
diff
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:
18110
diff
changeset
|
215 |
|
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
216 |
text {* Parsing Case expressions *} |
18097 | 217 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
218 |
syntax |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
219 |
"_pat" :: "'a" |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
220 |
"_variable" :: "'a" |
26046 | 221 |
"_noargs" :: "'a" |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
222 |
|
18097 | 223 |
translations |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
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:
29141
diff
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:
29141
diff
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:
18110
diff
changeset
|
227 |
|
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
228 |
parse_translation {* |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
229 |
(* rewrites (_pat x) => (return) *) |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
230 |
(* rewrites (_variable x t) => (Abs_CFun (%x. t)) *) |
19104 | 231 |
[("_pat", K (Syntax.const "Fixrec.return")), |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
232 |
mk_binder_tr ("_variable", "Abs_CFun")]; |
18112
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
233 |
*} |
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
234 |
|
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
235 |
text {* Printing Case expressions *} |
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
236 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
237 |
syntax |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
238 |
"_match" :: "'a" |
18097 | 239 |
|
18112
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
240 |
print_translation {* |
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
241 |
let |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
242 |
fun dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax unit_when},_) $ t) = |
26046 | 243 |
(Syntax.const "_noargs", t) |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
18112
diff
changeset
|
245 |
let |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
246 |
val (v1, t1) = dest_LAM t; |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
247 |
val (v2, t2) = dest_LAM t1; |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
248 |
in (Syntax.const "_args" $ v1 $ v2, t2) end |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
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:
18110
diff
changeset
|
250 |
let |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
251 |
val abs = case t of Abs abs => abs |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
252 |
| _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0); |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
253 |
val (x, t') = atomic_abs_tr' abs; |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
254 |
in (Syntax.const "_variable" $ x, t') end |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
255 |
| dest_LAM _ = raise Match; (* too few vars: abort translation *) |
18097 | 256 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
257 |
fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] = |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
258 |
let val (v, t) = dest_LAM r; |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
259 |
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:
18110
diff
changeset
|
260 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
261 |
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:
18110
diff
changeset
|
262 |
*} |
18097 | 263 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
264 |
translations |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
265 |
"x" <= "_match Fixrec.return (_variable x)" |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
266 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
267 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
268 |
subsection {* Pattern combinators for data constructors *} |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
269 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
270 |
types ('a, 'b) pat = "'a \<rightarrow> 'b maybe" |
18097 | 271 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
272 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
273 |
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:
23152
diff
changeset
|
274 |
"cpair_pat p1 p2 = (\<Lambda>\<langle>x, y\<rangle>. |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
275 |
bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>\<langle>a, b\<rangle>)))" |
18097 | 276 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
277 |
definition |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
278 |
spair_pat :: |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
279 |
"('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:
23152
diff
changeset
|
280 |
"spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>\<langle>x, y\<rangle>)" |
18097 | 281 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
282 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
283 |
sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
284 |
"sinl_pat p = sscase\<cdot>p\<cdot>(\<Lambda> x. fail)" |
18097 | 285 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
286 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
287 |
sinr_pat :: "('b, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
288 |
"sinr_pat p = sscase\<cdot>(\<Lambda> x. fail)\<cdot>p" |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
289 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
290 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
291 |
up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
292 |
"up_pat p = fup\<cdot>p" |
18097 | 293 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
294 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
295 |
TT_pat :: "(tr, unit) pat" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
296 |
"TT_pat = (\<Lambda> b. If b then return\<cdot>() else fail fi)" |
18097 | 297 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
298 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
299 |
FF_pat :: "(tr, unit) pat" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
300 |
"FF_pat = (\<Lambda> b. If b then fail else return\<cdot>() fi)" |
18097 | 301 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
302 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
303 |
ONE_pat :: "(one, unit) pat" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
304 |
"ONE_pat = (\<Lambda> ONE. return\<cdot>())" |
18097 | 305 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
306 |
text {* Parse translations (patterns) *} |
18097 | 307 |
translations |
26046 | 308 |
"_pat (XCONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)" |
309 |
"_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)" |
|
310 |
"_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)" |
|
311 |
"_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)" |
|
312 |
"_pat (XCONST up\<cdot>x)" => "CONST up_pat (_pat x)" |
|
313 |
"_pat (XCONST TT)" => "CONST TT_pat" |
|
314 |
"_pat (XCONST FF)" => "CONST FF_pat" |
|
315 |
"_pat (XCONST ONE)" => "CONST ONE_pat" |
|
316 |
||
317 |
text {* CONST version is also needed for constructors with special syntax *} |
|
318 |
translations |
|
319 |
"_pat (CONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)" |
|
320 |
"_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:
18112
diff
changeset
|
321 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
322 |
text {* Parse translations (variables) *} |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
323 |
translations |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
324 |
"_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:
29141
diff
changeset
|
325 |
"_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:
29141
diff
changeset
|
326 |
"_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:
29141
diff
changeset
|
327 |
"_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:
29141
diff
changeset
|
328 |
"_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:
29141
diff
changeset
|
329 |
"_variable (XCONST TT) r" => "_variable _noargs r" |
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
330 |
"_variable (XCONST FF) r" => "_variable _noargs r" |
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
331 |
"_variable (XCONST ONE) r" => "_variable _noargs r" |
26046 | 332 |
|
333 |
translations |
|
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
334 |
"_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:
29141
diff
changeset
|
335 |
"_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r" |
18097 | 336 |
|
18112
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
337 |
text {* Print translations *} |
18097 | 338 |
translations |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
339 |
"CONST cpair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
340 |
<= "_match (CONST cpair_pat p1 p2) (_args v1 v2)" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
341 |
"CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
342 |
<= "_match (CONST spair_pat p1 p2) (_args v1 v2)" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
343 |
"CONST sinl\<cdot>(_match p1 v1)" <= "_match (CONST sinl_pat p1) v1" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
344 |
"CONST sinr\<cdot>(_match p1 v1)" <= "_match (CONST sinr_pat p1) v1" |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
345 |
"CONST up\<cdot>(_match p1 v1)" <= "_match (CONST up_pat p1) v1" |
26046 | 346 |
"CONST TT" <= "_match (CONST TT_pat) _noargs" |
347 |
"CONST FF" <= "_match (CONST FF_pat) _noargs" |
|
348 |
"CONST ONE" <= "_match (CONST ONE_pat) _noargs" |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
349 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
350 |
lemma cpair_pat1: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
351 |
"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:
18112
diff
changeset
|
352 |
apply (simp add: branch_def cpair_pat_def) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
353 |
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
354 |
done |
18097 | 355 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
356 |
lemma cpair_pat2: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
357 |
"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:
18112
diff
changeset
|
358 |
apply (simp add: branch_def cpair_pat_def) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
359 |
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
360 |
done |
18097 | 361 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
362 |
lemma cpair_pat3: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
363 |
"branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow> |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
364 |
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:
18112
diff
changeset
|
365 |
apply (simp add: branch_def cpair_pat_def) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
366 |
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
367 |
apply (rule_tac p="q\<cdot>y" in maybeE, simp_all) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
368 |
done |
18097 | 369 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
370 |
lemmas cpair_pat [simp] = |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
371 |
cpair_pat1 cpair_pat2 cpair_pat3 |
18097 | 372 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
373 |
lemma spair_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
374 |
"branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
375 |
"\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
376 |
\<Longrightarrow> branch (spair_pat p1 p2)\<cdot>r\<cdot>(:x, y:) = |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
377 |
branch (cpair_pat p1 p2)\<cdot>r\<cdot>\<langle>x, y\<rangle>" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
378 |
by (simp_all add: branch_def spair_pat_def) |
18097 | 379 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
380 |
lemma sinl_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
381 |
"branch (sinl_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
382 |
"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:
18112
diff
changeset
|
383 |
"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:
18112
diff
changeset
|
384 |
by (simp_all add: branch_def sinl_pat_def) |
18097 | 385 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
386 |
lemma sinr_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
387 |
"branch (sinr_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
388 |
"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:
18112
diff
changeset
|
389 |
"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:
18112
diff
changeset
|
390 |
by (simp_all add: branch_def sinr_pat_def) |
18097 | 391 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
392 |
lemma up_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
393 |
"branch (up_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
394 |
"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:
18112
diff
changeset
|
395 |
by (simp_all add: branch_def up_pat_def) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
396 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
397 |
lemma TT_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
398 |
"branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
399 |
"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:
18112
diff
changeset
|
400 |
"branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
401 |
by (simp_all add: branch_def TT_pat_def) |
18097 | 402 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
403 |
lemma FF_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
404 |
"branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
405 |
"branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
406 |
"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:
18112
diff
changeset
|
407 |
by (simp_all add: branch_def FF_pat_def) |
18097 | 408 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
409 |
lemma ONE_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
410 |
"branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>" |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
411 |
"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:
18112
diff
changeset
|
412 |
by (simp_all add: branch_def ONE_pat_def) |
18097 | 413 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
414 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
415 |
subsection {* Wildcards, as-patterns, and lazy patterns *} |
18112
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
416 |
|
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
417 |
syntax |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
418 |
"_as_pat" :: "[idt, 'a] \<Rightarrow> 'a" (infixr "\<as>" 10) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
419 |
"_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:
18110
diff
changeset
|
420 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
421 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
422 |
wild_pat :: "'a \<rightarrow> unit maybe" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
423 |
"wild_pat = (\<Lambda> x. return\<cdot>())" |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
424 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
425 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
426 |
as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
427 |
"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:
18110
diff
changeset
|
428 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
429 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
430 |
lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where |
28891 | 431 |
"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:
18112
diff
changeset
|
432 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
433 |
text {* Parse translations (patterns) *} |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
434 |
translations |
26046 | 435 |
"_pat _" => "CONST wild_pat" |
436 |
"_pat (_as_pat x y)" => "CONST as_pat (_pat y)" |
|
437 |
"_pat (_lazy_pat x)" => "CONST lazy_pat (_pat x)" |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
438 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
439 |
text {* Parse translations (variables) *} |
18112
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
440 |
translations |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
441 |
"_variable _ r" => "_variable _noargs r" |
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
442 |
"_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:
29141
diff
changeset
|
443 |
"_variable (_lazy_pat x) r" => "_variable x r" |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
444 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
445 |
text {* Print translations *} |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
446 |
translations |
26046 | 447 |
"_" <= "_match (CONST wild_pat) _noargs" |
29322
ae6f2b1559fa
renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents:
29141
diff
changeset
|
448 |
"_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_variable x) v)" |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
449 |
"_lazy_pat (_match p v)" <= "_match (CONST lazy_pat p) v" |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
450 |
|
19327 | 451 |
text {* Lazy patterns in lambda abstractions *} |
452 |
translations |
|
28891 | 453 |
"_cabs (_lazy_pat p) r" == "CONST Fixrec.cases oo (_Case1 (_lazy_pat p) r)" |
19327 | 454 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
455 |
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:
18112
diff
changeset
|
456 |
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:
18110
diff
changeset
|
457 |
|
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
458 |
lemma as_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
459 |
"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:
18112
diff
changeset
|
460 |
apply (simp add: branch_def as_pat_def) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
461 |
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
462 |
done |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
463 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
464 |
lemma lazy_pat [simp]: |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
465 |
"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:
18112
diff
changeset
|
466 |
"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:
18112
diff
changeset
|
467 |
"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:
18112
diff
changeset
|
468 |
apply (simp_all add: branch_def lazy_pat_def) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
469 |
apply (rule_tac [!] p="p\<cdot>x" in maybeE, simp_all) |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
470 |
done |
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
471 |
|
18112
dc1d6f588204
reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents:
18110
diff
changeset
|
472 |
|
16221 | 473 |
subsection {* Match functions for built-in types *} |
474 |
||
16776
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
475 |
defaultsort pcpo |
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
476 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
477 |
definition |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
478 |
match_UU :: "'a \<rightarrow> 'c maybe \<rightarrow> 'c maybe" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
479 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
480 |
"match_UU = strictify\<cdot>(\<Lambda> x k. fail)" |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
481 |
|
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
482 |
definition |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
483 |
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:
30131
diff
changeset
|
484 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
485 |
"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:
16758
diff
changeset
|
486 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
487 |
definition |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
488 |
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:
30131
diff
changeset
|
489 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
490 |
"match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)" |
16221 | 491 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
492 |
definition |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
493 |
match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
494 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
495 |
"match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)" |
16551 | 496 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
497 |
definition |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
498 |
match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
499 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
500 |
"match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)" |
16551 | 501 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
502 |
definition |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
503 |
match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
504 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
505 |
"match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)" |
16221 | 506 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
507 |
definition |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
508 |
match_ONE :: "one \<rightarrow> 'c maybe \<rightarrow> 'c maybe" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
509 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
510 |
"match_ONE = (\<Lambda> ONE k. k)" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
511 |
|
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
512 |
definition |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
513 |
match_TT :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
514 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
515 |
"match_TT = (\<Lambda> x k. If x then k else fail fi)" |
18094 | 516 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
517 |
definition |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
518 |
match_FF :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
519 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
520 |
"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:
16417
diff
changeset
|
521 |
|
16776
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
522 |
lemma match_UU_simps [simp]: |
31008
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
huffman
parents:
30914
diff
changeset
|
523 |
"match_UU\<cdot>\<bottom>\<cdot>k = \<bottom>" |
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
huffman
parents:
30914
diff
changeset
|
524 |
"x \<noteq> \<bottom> \<Longrightarrow> match_UU\<cdot>x\<cdot>k = fail" |
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
huffman
parents:
30914
diff
changeset
|
525 |
by (simp_all add: match_UU_def) |
16776
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
526 |
|
16221 | 527 |
lemma match_cpair_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
528 |
"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:
31738
diff
changeset
|
529 |
"match_cpair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y" |
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
31738
diff
changeset
|
530 |
by (simp_all add: match_cpair_def) |
16221 | 531 |
|
16551 | 532 |
lemma match_spair_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
533 |
"\<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:
30131
diff
changeset
|
534 |
"match_spair\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 535 |
by (simp_all add: match_spair_def) |
536 |
||
537 |
lemma match_sinl_simps [simp]: |
|
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
538 |
"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:
30912
diff
changeset
|
539 |
"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:
30131
diff
changeset
|
540 |
"match_sinl\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 541 |
by (simp_all add: match_sinl_def) |
542 |
||
543 |
lemma match_sinr_simps [simp]: |
|
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
544 |
"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:
30912
diff
changeset
|
545 |
"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:
30131
diff
changeset
|
546 |
"match_sinr\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 547 |
by (simp_all add: match_sinr_def) |
548 |
||
16221 | 549 |
lemma match_up_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
550 |
"match_up\<cdot>(up\<cdot>x)\<cdot>k = k\<cdot>x" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
551 |
"match_up\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16221 | 552 |
by (simp_all add: match_up_def) |
553 |
||
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
554 |
lemma match_ONE_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
555 |
"match_ONE\<cdot>ONE\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
556 |
"match_ONE\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 557 |
by (simp_all add: match_ONE_def) |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
558 |
|
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
559 |
lemma match_TT_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
560 |
"match_TT\<cdot>TT\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
561 |
"match_TT\<cdot>FF\<cdot>k = fail" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
562 |
"match_TT\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 563 |
by (simp_all add: match_TT_def) |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
564 |
|
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
565 |
lemma match_FF_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
566 |
"match_FF\<cdot>FF\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
567 |
"match_FF\<cdot>TT\<cdot>k = fail" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
568 |
"match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 569 |
by (simp_all add: match_FF_def) |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
570 |
|
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
571 |
subsection {* Mutual recursion *} |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
572 |
|
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
573 |
text {* |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
574 |
The following rules are used to prove unfolding theorems from |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
575 |
fixed-point definitions of mutually recursive functions. |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
576 |
*} |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
577 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
578 |
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:
31008
diff
changeset
|
579 |
by simp |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
580 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
581 |
lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'" |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
582 |
by simp |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
583 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
584 |
lemma Pair_eqD2: "(x, y) = (x', y') \<Longrightarrow> y = y'" |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
585 |
by simp |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
586 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
587 |
lemma def_cont_fix_eq: |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
588 |
"\<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:
31008
diff
changeset
|
589 |
by (simp, subst fix_eq, simp) |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
590 |
|
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
591 |
lemma def_cont_fix_ind: |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
592 |
"\<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:
31008
diff
changeset
|
593 |
by (simp add: fix_ind) |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
594 |
|
16463
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
595 |
text {* lemma for proving rewrite rules *} |
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
596 |
|
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
597 |
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:
16460
diff
changeset
|
598 |
by simp |
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
599 |
|
16221 | 600 |
|
16758 | 601 |
subsection {* Initializing the fixrec package *} |
16221 | 602 |
|
31738
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31095
diff
changeset
|
603 |
use "Tools/fixrec.ML" |
16221 | 604 |
|
31738
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31095
diff
changeset
|
605 |
setup {* Fixrec.setup *} |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
606 |
|
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
607 |
setup {* |
31738
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31095
diff
changeset
|
608 |
Fixrec.add_matchers |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
609 |
[ (@{const_name up}, @{const_name match_up}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
610 |
(@{const_name sinl}, @{const_name match_sinl}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
611 |
(@{const_name sinr}, @{const_name match_sinr}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
612 |
(@{const_name spair}, @{const_name match_spair}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
613 |
(@{const_name cpair}, @{const_name match_cpair}), |
33401
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
31738
diff
changeset
|
614 |
(@{const_name Pair}, @{const_name match_cpair}), |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
615 |
(@{const_name ONE}, @{const_name match_ONE}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
616 |
(@{const_name TT}, @{const_name match_TT}), |
31008
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
huffman
parents:
30914
diff
changeset
|
617 |
(@{const_name FF}, @{const_name match_FF}), |
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
huffman
parents:
30914
diff
changeset
|
618 |
(@{const_name UU}, @{const_name match_UU}) ] |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
619 |
*} |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
620 |
|
28891 | 621 |
hide (open) const return bind fail run cases |
19104 | 622 |
|
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
623 |
lemmas [fixrec_simp] = |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
624 |
run_strict run_fail run_return |
33429 | 625 |
mplus_strict mplus_fail mplus_return |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
626 |
spair_strict_iff |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
627 |
sinl_defined_iff |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
628 |
sinr_defined_iff |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
629 |
up_defined |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
630 |
ONE_defined |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
631 |
dist_eq_tr(1,2) |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
632 |
match_UU_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
633 |
match_cpair_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
634 |
match_spair_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
635 |
match_sinl_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
636 |
match_sinr_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
637 |
match_up_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
638 |
match_ONE_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
639 |
match_TT_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
640 |
match_FF_simps |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
641 |
|
16221 | 642 |
end |