author | haftmann |
Thu, 19 Jun 2025 17:15:40 +0200 | |
changeset 82734 | 89347c0cc6a3 |
parent 81583 | b6df83045178 |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/Fixrec.thy |
81574 | 2 |
Author: Franz Regensburger |
16221 | 3 |
Author: Amber Telfer and Brian Huffman |
4 |
*) |
|
5 |
||
6 |
theory Fixrec |
|
81574 | 7 |
imports Cprod Sprod Ssum Up One Tr Cfun |
69913 | 8 |
keywords "fixrec" :: thy_defn |
16221 | 9 |
begin |
10 |
||
81574 | 11 |
section \<open>Fixed point operator and admissibility\<close> |
12 |
||
13 |
subsection \<open>Iteration\<close> |
|
14 |
||
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
15 |
primrec iterate :: "nat \<Rightarrow> ('a \<rightarrow> 'a) \<rightarrow> ('a \<rightarrow> 'a)" |
81574 | 16 |
where |
17 |
"iterate 0 = (\<Lambda> F x. x)" |
|
18 |
| "iterate (Suc n) = (\<Lambda> F x. F\<cdot>(iterate n\<cdot>F\<cdot>x))" |
|
19 |
||
20 |
text \<open>Derive inductive properties of iterate from primitive recursion\<close> |
|
21 |
||
22 |
lemma iterate_0 [simp]: "iterate 0\<cdot>F\<cdot>x = x" |
|
23 |
by simp |
|
24 |
||
25 |
lemma iterate_Suc [simp]: "iterate (Suc n)\<cdot>F\<cdot>x = F\<cdot>(iterate n\<cdot>F\<cdot>x)" |
|
26 |
by simp |
|
27 |
||
28 |
declare iterate.simps [simp del] |
|
29 |
||
30 |
lemma iterate_Suc2: "iterate (Suc n)\<cdot>F\<cdot>x = iterate n\<cdot>F\<cdot>(F\<cdot>x)" |
|
31 |
by (induct n) simp_all |
|
32 |
||
33 |
lemma iterate_iterate: "iterate m\<cdot>F\<cdot>(iterate n\<cdot>F\<cdot>x) = iterate (m + n)\<cdot>F\<cdot>x" |
|
34 |
by (induct m) simp_all |
|
35 |
||
36 |
text \<open>The sequence of function iterations is a chain.\<close> |
|
37 |
||
38 |
lemma chain_iterate [simp]: "chain (\<lambda>i. iterate i\<cdot>F\<cdot>\<bottom>)" |
|
39 |
by (rule chainI, unfold iterate_Suc2, rule monofun_cfun_arg, rule minimal) |
|
40 |
||
41 |
||
42 |
subsection \<open>Least fixed point operator\<close> |
|
43 |
||
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
44 |
definition "fix" :: "('a::pcpo \<rightarrow> 'a) \<rightarrow> 'a" |
81574 | 45 |
where "fix = (\<Lambda> F. \<Squnion>i. iterate i\<cdot>F\<cdot>\<bottom>)" |
46 |
||
47 |
text \<open>Binder syntax for \<^term>\<open>fix\<close>\<close> |
|
48 |
||
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
49 |
abbreviation fix_syn :: "('a::pcpo \<Rightarrow> 'a) \<Rightarrow> 'a" (binder \<open>\<mu> \<close> 10) |
81574 | 50 |
where "fix_syn (\<lambda>x. f x) \<equiv> fix\<cdot>(\<Lambda> x. f x)" |
51 |
||
52 |
notation (ASCII) |
|
53 |
fix_syn (binder \<open>FIX \<close> 10) |
|
54 |
||
55 |
text \<open>Properties of \<^term>\<open>fix\<close>\<close> |
|
56 |
||
57 |
text \<open>direct connection between \<^term>\<open>fix\<close> and iteration\<close> |
|
58 |
||
59 |
lemma fix_def2: "fix\<cdot>F = (\<Squnion>i. iterate i\<cdot>F\<cdot>\<bottom>)" |
|
60 |
by (simp add: fix_def) |
|
61 |
||
62 |
lemma iterate_below_fix: "iterate n\<cdot>f\<cdot>\<bottom> \<sqsubseteq> fix\<cdot>f" |
|
63 |
unfolding fix_def2 |
|
64 |
using chain_iterate by (rule is_ub_thelub) |
|
65 |
||
66 |
text \<open> |
|
67 |
Kleene's fixed point theorems for continuous functions in pointed |
|
68 |
omega cpo's |
|
69 |
\<close> |
|
70 |
||
71 |
lemma fix_eq: "fix\<cdot>F = F\<cdot>(fix\<cdot>F)" |
|
72 |
apply (simp add: fix_def2) |
|
73 |
apply (subst lub_range_shift [of _ 1, symmetric]) |
|
74 |
apply (rule chain_iterate) |
|
75 |
apply (subst contlub_cfun_arg) |
|
76 |
apply (rule chain_iterate) |
|
77 |
apply simp |
|
78 |
done |
|
79 |
||
80 |
lemma fix_least_below: "F\<cdot>x \<sqsubseteq> x \<Longrightarrow> fix\<cdot>F \<sqsubseteq> x" |
|
81 |
apply (simp add: fix_def2) |
|
82 |
apply (rule lub_below) |
|
83 |
apply (rule chain_iterate) |
|
84 |
apply (induct_tac i) |
|
85 |
apply simp |
|
86 |
apply simp |
|
87 |
apply (erule rev_below_trans) |
|
88 |
apply (erule monofun_cfun_arg) |
|
89 |
done |
|
90 |
||
91 |
lemma fix_least: "F\<cdot>x = x \<Longrightarrow> fix\<cdot>F \<sqsubseteq> x" |
|
92 |
by (rule fix_least_below) simp |
|
93 |
||
94 |
lemma fix_eqI: |
|
95 |
assumes fixed: "F\<cdot>x = x" |
|
96 |
and least: "\<And>z. F\<cdot>z = z \<Longrightarrow> x \<sqsubseteq> z" |
|
97 |
shows "fix\<cdot>F = x" |
|
98 |
apply (rule below_antisym) |
|
99 |
apply (rule fix_least [OF fixed]) |
|
100 |
apply (rule least [OF fix_eq [symmetric]]) |
|
101 |
done |
|
102 |
||
103 |
lemma fix_eq2: "f \<equiv> fix\<cdot>F \<Longrightarrow> f = F\<cdot>f" |
|
104 |
by (simp add: fix_eq [symmetric]) |
|
105 |
||
106 |
lemma fix_eq3: "f \<equiv> fix\<cdot>F \<Longrightarrow> f\<cdot>x = F\<cdot>f\<cdot>x" |
|
107 |
by (erule fix_eq2 [THEN cfun_fun_cong]) |
|
108 |
||
109 |
lemma fix_eq4: "f = fix\<cdot>F \<Longrightarrow> f = F\<cdot>f" |
|
110 |
by (erule ssubst) (rule fix_eq) |
|
111 |
||
112 |
lemma fix_eq5: "f = fix\<cdot>F \<Longrightarrow> f\<cdot>x = F\<cdot>f\<cdot>x" |
|
113 |
by (erule fix_eq4 [THEN cfun_fun_cong]) |
|
114 |
||
115 |
text \<open>strictness of \<^term>\<open>fix\<close>\<close> |
|
116 |
||
117 |
lemma fix_bottom_iff: "fix\<cdot>F = \<bottom> \<longleftrightarrow> F\<cdot>\<bottom> = \<bottom>" |
|
118 |
apply (rule iffI) |
|
119 |
apply (erule subst) |
|
120 |
apply (rule fix_eq [symmetric]) |
|
121 |
apply (erule fix_least [THEN bottomI]) |
|
122 |
done |
|
123 |
||
124 |
lemma fix_strict: "F\<cdot>\<bottom> = \<bottom> \<Longrightarrow> fix\<cdot>F = \<bottom>" |
|
125 |
by (simp add: fix_bottom_iff) |
|
126 |
||
127 |
lemma fix_defined: "F\<cdot>\<bottom> \<noteq> \<bottom> \<Longrightarrow> fix\<cdot>F \<noteq> \<bottom>" |
|
128 |
by (simp add: fix_bottom_iff) |
|
129 |
||
130 |
text \<open>\<^term>\<open>fix\<close> applied to identity and constant functions\<close> |
|
131 |
||
132 |
lemma fix_id: "(\<mu> x. x) = \<bottom>" |
|
133 |
by (simp add: fix_strict) |
|
134 |
||
135 |
lemma fix_const: "(\<mu> x. c) = c" |
|
136 |
by (subst fix_eq) simp |
|
137 |
||
138 |
||
139 |
subsection \<open>Fixed point induction\<close> |
|
140 |
||
141 |
lemma fix_ind: "adm P \<Longrightarrow> P \<bottom> \<Longrightarrow> (\<And>x. P x \<Longrightarrow> P (F\<cdot>x)) \<Longrightarrow> P (fix\<cdot>F)" |
|
142 |
unfolding fix_def2 |
|
143 |
apply (erule admD) |
|
144 |
apply (rule chain_iterate) |
|
145 |
apply (rule nat_induct, simp_all) |
|
146 |
done |
|
147 |
||
148 |
lemma cont_fix_ind: "cont F \<Longrightarrow> adm P \<Longrightarrow> P \<bottom> \<Longrightarrow> (\<And>x. P x \<Longrightarrow> P (F x)) \<Longrightarrow> P (fix\<cdot>(Abs_cfun F))" |
|
149 |
by (simp add: fix_ind) |
|
150 |
||
151 |
lemma def_fix_ind: "\<lbrakk>f \<equiv> fix\<cdot>F; adm P; P \<bottom>; \<And>x. P x \<Longrightarrow> P (F\<cdot>x)\<rbrakk> \<Longrightarrow> P f" |
|
152 |
by (simp add: fix_ind) |
|
153 |
||
154 |
lemma fix_ind2: |
|
155 |
assumes adm: "adm P" |
|
156 |
assumes 0: "P \<bottom>" and 1: "P (F\<cdot>\<bottom>)" |
|
157 |
assumes step: "\<And>x. \<lbrakk>P x; P (F\<cdot>x)\<rbrakk> \<Longrightarrow> P (F\<cdot>(F\<cdot>x))" |
|
158 |
shows "P (fix\<cdot>F)" |
|
159 |
unfolding fix_def2 |
|
160 |
apply (rule admD [OF adm chain_iterate]) |
|
161 |
apply (rule nat_less_induct) |
|
162 |
apply (case_tac n) |
|
163 |
apply (simp add: 0) |
|
164 |
apply (case_tac nat) |
|
165 |
apply (simp add: 1) |
|
166 |
apply (frule_tac x=nat in spec) |
|
167 |
apply (simp add: step) |
|
168 |
done |
|
169 |
||
170 |
lemma parallel_fix_ind: |
|
171 |
assumes adm: "adm (\<lambda>x. P (fst x) (snd x))" |
|
172 |
assumes base: "P \<bottom> \<bottom>" |
|
173 |
assumes step: "\<And>x y. P x y \<Longrightarrow> P (F\<cdot>x) (G\<cdot>y)" |
|
174 |
shows "P (fix\<cdot>F) (fix\<cdot>G)" |
|
175 |
proof - |
|
176 |
from adm have adm': "adm (case_prod P)" |
|
177 |
unfolding split_def . |
|
178 |
have "P (iterate i\<cdot>F\<cdot>\<bottom>) (iterate i\<cdot>G\<cdot>\<bottom>)" for i |
|
179 |
by (induct i) (simp add: base, simp add: step) |
|
180 |
then have "\<And>i. case_prod P (iterate i\<cdot>F\<cdot>\<bottom>, iterate i\<cdot>G\<cdot>\<bottom>)" |
|
181 |
by simp |
|
182 |
then have "case_prod P (\<Squnion>i. (iterate i\<cdot>F\<cdot>\<bottom>, iterate i\<cdot>G\<cdot>\<bottom>))" |
|
183 |
by - (rule admD [OF adm'], simp, assumption) |
|
184 |
then have "case_prod P (\<Squnion>i. iterate i\<cdot>F\<cdot>\<bottom>, \<Squnion>i. iterate i\<cdot>G\<cdot>\<bottom>)" |
|
185 |
by (simp add: lub_Pair) |
|
186 |
then have "P (\<Squnion>i. iterate i\<cdot>F\<cdot>\<bottom>) (\<Squnion>i. iterate i\<cdot>G\<cdot>\<bottom>)" |
|
187 |
by simp |
|
188 |
then show "P (fix\<cdot>F) (fix\<cdot>G)" |
|
189 |
by (simp add: fix_def2) |
|
190 |
qed |
|
191 |
||
192 |
lemma cont_parallel_fix_ind: |
|
193 |
assumes "cont F" and "cont G" |
|
194 |
assumes "adm (\<lambda>x. P (fst x) (snd x))" |
|
195 |
assumes "P \<bottom> \<bottom>" |
|
196 |
assumes "\<And>x y. P x y \<Longrightarrow> P (F x) (G y)" |
|
197 |
shows "P (fix\<cdot>(Abs_cfun F)) (fix\<cdot>(Abs_cfun G))" |
|
198 |
by (rule parallel_fix_ind) (simp_all add: assms) |
|
199 |
||
200 |
||
201 |
subsection \<open>Fixed-points on product types\<close> |
|
202 |
||
203 |
text \<open> |
|
204 |
Bekic's Theorem: Simultaneous fixed points over pairs |
|
205 |
can be written in terms of separate fixed points. |
|
206 |
\<close> |
|
207 |
||
208 |
lemma fix_cprod: |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
209 |
fixes F :: "'a::pcpo \<times> 'b::pcpo \<rightarrow> 'a \<times> 'b" |
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
210 |
shows |
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
211 |
"fix\<cdot>F = |
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
212 |
(\<mu> x. fst (F\<cdot>(x, \<mu> y. snd (F\<cdot>(x, y)))), |
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
213 |
\<mu> y. snd (F\<cdot>(\<mu> x. fst (F\<cdot>(x, \<mu> y. snd (F\<cdot>(x, y)))), y)))" |
81574 | 214 |
(is "fix\<cdot>F = (?x, ?y)") |
215 |
proof (rule fix_eqI) |
|
216 |
have *: "fst (F\<cdot>(?x, ?y)) = ?x" |
|
217 |
by (rule trans [symmetric, OF fix_eq], simp) |
|
218 |
have "snd (F\<cdot>(?x, ?y)) = ?y" |
|
219 |
by (rule trans [symmetric, OF fix_eq], simp) |
|
220 |
with * show "F\<cdot>(?x, ?y) = (?x, ?y)" |
|
221 |
by (simp add: prod_eq_iff) |
|
222 |
next |
|
223 |
fix z |
|
224 |
assume F_z: "F\<cdot>z = z" |
|
225 |
obtain x y where z: "z = (x, y)" by (rule prod.exhaust) |
|
226 |
from F_z z have F_x: "fst (F\<cdot>(x, y)) = x" by simp |
|
227 |
from F_z z have F_y: "snd (F\<cdot>(x, y)) = y" by simp |
|
228 |
let ?y1 = "\<mu> y. snd (F\<cdot>(x, y))" |
|
229 |
have "?y1 \<sqsubseteq> y" |
|
230 |
by (rule fix_least) (simp add: F_y) |
|
231 |
then have "fst (F\<cdot>(x, ?y1)) \<sqsubseteq> fst (F\<cdot>(x, y))" |
|
232 |
by (simp add: fst_monofun monofun_cfun) |
|
233 |
with F_x have "fst (F\<cdot>(x, ?y1)) \<sqsubseteq> x" |
|
234 |
by simp |
|
235 |
then have *: "?x \<sqsubseteq> x" |
|
236 |
by (simp add: fix_least_below) |
|
237 |
then have "snd (F\<cdot>(?x, y)) \<sqsubseteq> snd (F\<cdot>(x, y))" |
|
238 |
by (simp add: snd_monofun monofun_cfun) |
|
239 |
with F_y have "snd (F\<cdot>(?x, y)) \<sqsubseteq> y" |
|
240 |
by simp |
|
241 |
then have "?y \<sqsubseteq> y" |
|
242 |
by (simp add: fix_least_below) |
|
243 |
with z * show "(?x, ?y) \<sqsubseteq> z" |
|
244 |
by simp |
|
245 |
qed |
|
246 |
||
247 |
||
248 |
section "Package for defining recursive functions in HOLCF" |
|
249 |
||
62175 | 250 |
subsection \<open>Pattern-match monad\<close> |
16221 | 251 |
|
49759
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
huffman
parents:
48891
diff
changeset
|
252 |
pcpodef 'a match = "UNIV::(one ++ 'a u) set" |
29063
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents:
28891
diff
changeset
|
253 |
by simp_all |
16221 | 254 |
|
29141 | 255 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
256 |
fail :: "'a match" where |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
257 |
"fail = Abs_match (sinl\<cdot>ONE)" |
16221 | 258 |
|
29141 | 259 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
260 |
succeed :: "'a \<rightarrow> 'a match" where |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
261 |
"succeed = (\<Lambda> x. Abs_match (sinr\<cdot>(up\<cdot>x)))" |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
262 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
263 |
lemma matchE [case_names bottom fail succeed, cases type: match]: |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
264 |
"\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = succeed\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
265 |
unfolding fail_def succeed_def |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
266 |
apply (cases p, rename_tac r) |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
267 |
apply (rule_tac p=r in ssumE, simp add: Abs_match_strict) |
16221 | 268 |
apply (rule_tac p=x in oneE, simp, simp) |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
269 |
apply (rule_tac p=y in upE, simp, simp add: cont_Abs_match) |
16221 | 270 |
done |
271 |
||
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
272 |
lemma succeed_defined [simp]: "succeed\<cdot>x \<noteq> \<bottom>" |
41029
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
huffman
parents:
40834
diff
changeset
|
273 |
by (simp add: succeed_def cont_Abs_match Abs_match_bottom_iff) |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
274 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
275 |
lemma fail_defined [simp]: "fail \<noteq> \<bottom>" |
41029
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
huffman
parents:
40834
diff
changeset
|
276 |
by (simp add: fail_def Abs_match_bottom_iff) |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
277 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
278 |
lemma succeed_eq [simp]: "(succeed\<cdot>x = succeed\<cdot>y) = (x = y)" |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
279 |
by (simp add: succeed_def cont_Abs_match Abs_match_inject) |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
280 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
281 |
lemma succeed_neq_fail [simp]: |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
282 |
"succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x" |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
283 |
by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject) |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
284 |
|
81577 | 285 |
|
62175 | 286 |
subsubsection \<open>Run operator\<close> |
16221 | 287 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
288 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
289 |
run :: "'a match \<rightarrow> 'a::pcpo" where |
40735 | 290 |
"run = (\<Lambda> m. sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)\<cdot>(Rep_match m))" |
16221 | 291 |
|
62175 | 292 |
text \<open>rewrite rules for run\<close> |
16221 | 293 |
|
294 |
lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>" |
|
40735 | 295 |
unfolding run_def |
296 |
by (simp add: cont_Rep_match Rep_match_strict) |
|
16221 | 297 |
|
298 |
lemma run_fail [simp]: "run\<cdot>fail = \<bottom>" |
|
40735 | 299 |
unfolding run_def fail_def |
300 |
by (simp add: cont_Rep_match Abs_match_inverse) |
|
16221 | 301 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
302 |
lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>x) = x" |
40735 | 303 |
unfolding run_def succeed_def |
304 |
by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse) |
|
16221 | 305 |
|
81577 | 306 |
|
62175 | 307 |
subsubsection \<open>Monad plus operator\<close> |
16221 | 308 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
309 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
310 |
mplus :: "'a match \<rightarrow> 'a match \<rightarrow> 'a match" where |
40735 | 311 |
"mplus = (\<Lambda> m1 m2. sscase\<cdot>(\<Lambda> _. m2)\<cdot>(\<Lambda> _. m1)\<cdot>(Rep_match m1))" |
18097 | 312 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
313 |
abbreviation |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69913
diff
changeset
|
314 |
mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match" (infixr \<open>+++\<close> 65) where |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
315 |
"m1 +++ m2 == mplus\<cdot>m1\<cdot>m2" |
16221 | 316 |
|
62175 | 317 |
text \<open>rewrite rules for mplus\<close> |
16221 | 318 |
|
319 |
lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>" |
|
40735 | 320 |
unfolding mplus_def |
40834
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
huffman
parents:
40795
diff
changeset
|
321 |
by (simp add: cont_Rep_match Rep_match_strict) |
16221 | 322 |
|
323 |
lemma mplus_fail [simp]: "fail +++ m = m" |
|
40735 | 324 |
unfolding mplus_def fail_def |
40834
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
huffman
parents:
40795
diff
changeset
|
325 |
by (simp add: cont_Rep_match Abs_match_inverse) |
16221 | 326 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
327 |
lemma mplus_succeed [simp]: "succeed\<cdot>x +++ m = succeed\<cdot>x" |
40735 | 328 |
unfolding mplus_def succeed_def |
40834
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
huffman
parents:
40795
diff
changeset
|
329 |
by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse) |
16221 | 330 |
|
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
331 |
lemma mplus_fail2 [simp]: "m +++ fail = m" |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
332 |
by (cases m, simp_all) |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
333 |
|
16221 | 334 |
lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)" |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
335 |
by (cases x, simp_all) |
16221 | 336 |
|
81577 | 337 |
|
62175 | 338 |
subsection \<open>Match functions for built-in types\<close> |
16221 | 339 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
340 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
341 |
match_bottom :: "'a::pcpo \<rightarrow> 'c match \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
342 |
where |
40768 | 343 |
"match_bottom = (\<Lambda> x k. seq\<cdot>x\<cdot>fail)" |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
344 |
|
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
345 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
346 |
match_Pair :: "'a \<times> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
347 |
where |
39807 | 348 |
"match_Pair = (\<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
|
349 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
350 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
351 |
match_spair :: "'a::pcpo \<otimes> 'b::pcpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c::pcpo match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
352 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
353 |
"match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)" |
16221 | 354 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
355 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
356 |
match_sinl :: "'a::pcpo \<oplus> 'b::pcpo \<rightarrow> ('a \<rightarrow> 'c::pcpo match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
357 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
358 |
"match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)" |
16551 | 359 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
360 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
361 |
match_sinr :: "'a::pcpo \<oplus> 'b::pcpo \<rightarrow> ('b \<rightarrow> 'c::pcpo match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
362 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
363 |
"match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)" |
16551 | 364 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
365 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
366 |
match_up :: "'a u \<rightarrow> ('a \<rightarrow> 'c::pcpo match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
367 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
368 |
"match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)" |
16221 | 369 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
370 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
371 |
match_ONE :: "one \<rightarrow> 'c::pcpo match \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
372 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
373 |
"match_ONE = (\<Lambda> ONE k. k)" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
374 |
|
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
375 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
376 |
match_TT :: "tr \<rightarrow> 'c::pcpo match \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
377 |
where |
40322
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
huffman
parents:
40046
diff
changeset
|
378 |
"match_TT = (\<Lambda> x k. If x then k else fail)" |
65380
ae93953746fc
eliminated Plain_HOLCF.thy (see also 8e92772bc0e8): it was modeled after HOL/Plain.thy which was discontinued later;
wenzelm
parents:
62175
diff
changeset
|
379 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
380 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
381 |
match_FF :: "tr \<rightarrow> 'c::pcpo match \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
382 |
where |
40322
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
huffman
parents:
40046
diff
changeset
|
383 |
"match_FF = (\<Lambda> x k. If x then fail else k)" |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
384 |
|
40768 | 385 |
lemma match_bottom_simps [simp]: |
40795
c52cd8bc426d
change match_bottom_simps to produce if-then-else, making more uses of bottom-patterns work with fixrec
huffman
parents:
40774
diff
changeset
|
386 |
"match_bottom\<cdot>x\<cdot>k = (if x = \<bottom> then \<bottom> else fail)" |
c52cd8bc426d
change match_bottom_simps to produce if-then-else, making more uses of bottom-patterns work with fixrec
huffman
parents:
40774
diff
changeset
|
387 |
by (simp add: match_bottom_def) |
16776
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
388 |
|
39807 | 389 |
lemma match_Pair_simps [simp]: |
390 |
"match_Pair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y" |
|
391 |
by (simp_all add: match_Pair_def) |
|
16221 | 392 |
|
16551 | 393 |
lemma match_spair_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
394 |
"\<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
|
395 |
"match_spair\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 396 |
by (simp_all add: match_spair_def) |
397 |
||
398 |
lemma match_sinl_simps [simp]: |
|
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
399 |
"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
|
400 |
"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
|
401 |
"match_sinl\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 402 |
by (simp_all add: match_sinl_def) |
403 |
||
404 |
lemma match_sinr_simps [simp]: |
|
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
405 |
"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
|
406 |
"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
|
407 |
"match_sinr\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 408 |
by (simp_all add: match_sinr_def) |
409 |
||
16221 | 410 |
lemma match_up_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
411 |
"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
|
412 |
"match_up\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16221 | 413 |
by (simp_all add: match_up_def) |
414 |
||
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
415 |
lemma match_ONE_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
416 |
"match_ONE\<cdot>ONE\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
417 |
"match_ONE\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 418 |
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
|
419 |
|
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
420 |
lemma match_TT_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
421 |
"match_TT\<cdot>TT\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
422 |
"match_TT\<cdot>FF\<cdot>k = fail" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
423 |
"match_TT\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 424 |
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
|
425 |
|
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
426 |
lemma match_FF_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
427 |
"match_FF\<cdot>FF\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
428 |
"match_FF\<cdot>TT\<cdot>k = fail" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
429 |
"match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 430 |
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
|
431 |
|
81577 | 432 |
|
62175 | 433 |
subsection \<open>Mutual recursion\<close> |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
434 |
|
62175 | 435 |
text \<open> |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
436 |
The following rules are used to prove unfolding theorems from |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
437 |
fixed-point definitions of mutually recursive functions. |
62175 | 438 |
\<close> |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
439 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
440 |
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
|
441 |
by simp |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
442 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
443 |
lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'" |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
444 |
by simp |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
445 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
446 |
lemma Pair_eqD2: "(x, y) = (x', y') \<Longrightarrow> y = y'" |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
447 |
by simp |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
448 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
449 |
lemma def_cont_fix_eq: |
40327 | 450 |
"\<lbrakk>f \<equiv> fix\<cdot>(Abs_cfun F); cont F\<rbrakk> \<Longrightarrow> f = F f" |
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
451 |
by (simp, subst fix_eq, simp) |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
452 |
|
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
453 |
lemma def_cont_fix_ind: |
40327 | 454 |
"\<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" |
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
455 |
by (simp add: fix_ind) |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
456 |
|
62175 | 457 |
text \<open>lemma for proving rewrite rules\<close> |
16463
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
458 |
|
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
459 |
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
|
460 |
by simp |
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
461 |
|
16221 | 462 |
|
62175 | 463 |
subsection \<open>Initializing the fixrec package\<close> |
16221 | 464 |
|
69605 | 465 |
ML_file \<open>Tools/holcf_library.ML\<close> |
466 |
ML_file \<open>Tools/fixrec.ML\<close> |
|
16221 | 467 |
|
62175 | 468 |
method_setup fixrec_simp = \<open> |
47432 | 469 |
Scan.succeed (SIMPLE_METHOD' o Fixrec.fixrec_simp_tac) |
62175 | 470 |
\<close> "pattern prover for fixrec constants" |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
471 |
|
62175 | 472 |
setup \<open> |
31738
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31095
diff
changeset
|
473 |
Fixrec.add_matchers |
69597 | 474 |
[ (\<^const_name>\<open>up\<close>, \<^const_name>\<open>match_up\<close>), |
475 |
(\<^const_name>\<open>sinl\<close>, \<^const_name>\<open>match_sinl\<close>), |
|
476 |
(\<^const_name>\<open>sinr\<close>, \<^const_name>\<open>match_sinr\<close>), |
|
477 |
(\<^const_name>\<open>spair\<close>, \<^const_name>\<open>match_spair\<close>), |
|
478 |
(\<^const_name>\<open>Pair\<close>, \<^const_name>\<open>match_Pair\<close>), |
|
479 |
(\<^const_name>\<open>ONE\<close>, \<^const_name>\<open>match_ONE\<close>), |
|
480 |
(\<^const_name>\<open>TT\<close>, \<^const_name>\<open>match_TT\<close>), |
|
481 |
(\<^const_name>\<open>FF\<close>, \<^const_name>\<open>match_FF\<close>), |
|
482 |
(\<^const_name>\<open>bottom\<close>, \<^const_name>\<open>match_bottom\<close>) ] |
|
62175 | 483 |
\<close> |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
484 |
|
37109
e67760c1b851
move unused pattern match syntax stuff into HOLCF/ex
huffman
parents:
37108
diff
changeset
|
485 |
hide_const (open) succeed fail run |
19104 | 486 |
|
16221 | 487 |
end |