author | haftmann |
Fri, 27 Jun 2025 08:09:26 +0200 | |
changeset 82775 | 61c39a9e5415 |
parent 81733 | 94f018b2a4fb |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/Library/List_Cpo.thy |
39143 | 2 |
Author: Brian Huffman |
3 |
*) |
|
4 |
||
62175 | 5 |
section \<open>Lists as a complete partial order\<close> |
39143 | 6 |
|
7 |
theory List_Cpo |
|
8 |
imports HOLCF |
|
9 |
begin |
|
10 |
||
62175 | 11 |
subsection \<open>Lists are a partial order\<close> |
39143 | 12 |
|
13 |
instantiation list :: (po) po |
|
14 |
begin |
|
15 |
||
16 |
definition |
|
67399 | 17 |
"xs \<sqsubseteq> ys \<longleftrightarrow> list_all2 (\<sqsubseteq>) xs ys" |
39143 | 18 |
|
81733 | 19 |
instance |
20 |
proof |
|
21 |
fix xs ys zs :: "'a list" |
|
22 |
show "xs \<sqsubseteq> xs" |
|
23 |
using below_refl |
|
39143 | 24 |
unfolding below_list_def |
25 |
by (rule list_all2_refl) |
|
81733 | 26 |
show "xs \<sqsubseteq> ys \<Longrightarrow> ys \<sqsubseteq> zs \<Longrightarrow> xs \<sqsubseteq> zs" |
27 |
using below_trans |
|
39143 | 28 |
unfolding below_list_def |
29 |
by (rule list_all2_trans) |
|
81733 | 30 |
show "xs \<sqsubseteq> ys \<Longrightarrow> ys \<sqsubseteq> xs \<Longrightarrow> xs = ys" |
31 |
using below_antisym |
|
39143 | 32 |
unfolding below_list_def |
33 |
by (rule list_all2_antisym) |
|
34 |
qed |
|
35 |
||
36 |
end |
|
37 |
||
38 |
lemma below_list_simps [simp]: |
|
39 |
"[] \<sqsubseteq> []" |
|
40 |
"x # xs \<sqsubseteq> y # ys \<longleftrightarrow> x \<sqsubseteq> y \<and> xs \<sqsubseteq> ys" |
|
41 |
"\<not> [] \<sqsubseteq> y # ys" |
|
42 |
"\<not> x # xs \<sqsubseteq> []" |
|
43 |
by (simp_all add: below_list_def) |
|
44 |
||
45 |
lemma Nil_below_iff [simp]: "[] \<sqsubseteq> xs \<longleftrightarrow> xs = []" |
|
46 |
by (cases xs, simp_all) |
|
47 |
||
48 |
lemma below_Nil_iff [simp]: "xs \<sqsubseteq> [] \<longleftrightarrow> xs = []" |
|
49 |
by (cases xs, simp_all) |
|
50 |
||
39966 | 51 |
lemma list_below_induct [consumes 1, case_names Nil Cons]: |
52 |
assumes "xs \<sqsubseteq> ys" |
|
53 |
assumes 1: "P [] []" |
|
54 |
assumes 2: "\<And>x y xs ys. \<lbrakk>x \<sqsubseteq> y; xs \<sqsubseteq> ys; P xs ys\<rbrakk> \<Longrightarrow> P (x # xs) (y # ys)" |
|
55 |
shows "P xs ys" |
|
62175 | 56 |
using \<open>xs \<sqsubseteq> ys\<close> |
39966 | 57 |
proof (induct xs arbitrary: ys) |
58 |
case Nil thus ?case by (simp add: 1) |
|
59 |
next |
|
60 |
case (Cons x xs) thus ?case by (cases ys, simp_all add: 2) |
|
61 |
qed |
|
62 |
||
63 |
lemma list_below_cases: |
|
64 |
assumes "xs \<sqsubseteq> ys" |
|
65 |
obtains "xs = []" and "ys = []" | |
|
66 |
x y xs' ys' where "xs = x # xs'" and "ys = y # ys'" |
|
67 |
using assms by (cases xs, simp, cases ys, auto) |
|
68 |
||
39143 | 69 |
text "Thanks to Joachim Breitner" |
70 |
||
71 |
lemma list_Cons_below: |
|
72 |
assumes "a # as \<sqsubseteq> xs" |
|
73 |
obtains b and bs where "a \<sqsubseteq> b" and "as \<sqsubseteq> bs" and "xs = b # bs" |
|
74 |
using assms by (cases xs, auto) |
|
75 |
||
76 |
lemma list_below_Cons: |
|
77 |
assumes "xs \<sqsubseteq> b # bs" |
|
78 |
obtains a and as where "a \<sqsubseteq> b" and "as \<sqsubseteq> bs" and "xs = a # as" |
|
79 |
using assms by (cases xs, auto) |
|
80 |
||
81 |
lemma hd_mono: "xs \<sqsubseteq> ys \<Longrightarrow> hd xs \<sqsubseteq> hd ys" |
|
82 |
by (cases xs, simp, cases ys, simp, simp) |
|
83 |
||
84 |
lemma tl_mono: "xs \<sqsubseteq> ys \<Longrightarrow> tl xs \<sqsubseteq> tl ys" |
|
85 |
by (cases xs, simp, cases ys, simp, simp) |
|
86 |
||
87 |
lemma ch2ch_hd [simp]: "chain (\<lambda>i. S i) \<Longrightarrow> chain (\<lambda>i. hd (S i))" |
|
88 |
by (rule chainI, rule hd_mono, erule chainE) |
|
89 |
||
90 |
lemma ch2ch_tl [simp]: "chain (\<lambda>i. S i) \<Longrightarrow> chain (\<lambda>i. tl (S i))" |
|
91 |
by (rule chainI, rule tl_mono, erule chainE) |
|
92 |
||
93 |
lemma below_same_length: "xs \<sqsubseteq> ys \<Longrightarrow> length xs = length ys" |
|
94 |
unfolding below_list_def by (rule list_all2_lengthD) |
|
95 |
||
39966 | 96 |
lemma list_chain_induct [consumes 1, case_names Nil Cons]: |
97 |
assumes "chain S" |
|
98 |
assumes 1: "P (\<lambda>i. [])" |
|
99 |
assumes 2: "\<And>A B. chain A \<Longrightarrow> chain B \<Longrightarrow> P B \<Longrightarrow> P (\<lambda>i. A i # B i)" |
|
100 |
shows "P S" |
|
62175 | 101 |
using \<open>chain S\<close> |
39966 | 102 |
proof (induct "S 0" arbitrary: S) |
103 |
case Nil |
|
62175 | 104 |
have "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono [OF \<open>chain S\<close>]) |
39966 | 105 |
with Nil have "\<forall>i. S i = []" by simp |
106 |
thus ?case by (simp add: 1) |
|
107 |
next |
|
108 |
case (Cons x xs) |
|
62175 | 109 |
have "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono [OF \<open>chain S\<close>]) |
39966 | 110 |
hence *: "\<forall>i. S i \<noteq> []" by (rule all_forward, insert Cons) auto |
111 |
have "chain (\<lambda>i. hd (S i))" and "chain (\<lambda>i. tl (S i))" |
|
62175 | 112 |
using \<open>chain S\<close> by simp_all |
39966 | 113 |
moreover have "P (\<lambda>i. tl (S i))" |
62175 | 114 |
using \<open>chain S\<close> and \<open>x # xs = S 0\<close> [symmetric] |
39966 | 115 |
by (simp add: Cons(1)) |
116 |
ultimately have "P (\<lambda>i. hd (S i) # tl (S i))" |
|
117 |
by (rule 2) |
|
118 |
thus "P S" by (simp add: *) |
|
119 |
qed |
|
120 |
||
39143 | 121 |
lemma list_chain_cases: |
122 |
assumes S: "chain S" |
|
39966 | 123 |
obtains "S = (\<lambda>i. [])" | |
124 |
A B where "chain A" and "chain B" and "S = (\<lambda>i. A i # B i)" |
|
125 |
using S by (induct rule: list_chain_induct) simp_all |
|
39143 | 126 |
|
62175 | 127 |
subsection \<open>Lists are a complete partial order\<close> |
39143 | 128 |
|
129 |
lemma is_lub_Cons: |
|
130 |
assumes A: "range A <<| x" |
|
131 |
assumes B: "range B <<| xs" |
|
132 |
shows "range (\<lambda>i. A i # B i) <<| x # xs" |
|
133 |
using assms |
|
39968 | 134 |
unfolding is_lub_def is_ub_def |
39143 | 135 |
by (clarsimp, case_tac u, simp_all) |
136 |
||
137 |
instance list :: (cpo) cpo |
|
138 |
proof |
|
139 |
fix S :: "nat \<Rightarrow> 'a list" |
|
68780 | 140 |
assume "chain S" |
141 |
then show "\<exists>x. range S <<| x" |
|
39966 | 142 |
proof (induct rule: list_chain_induct) |
68780 | 143 |
case Nil |
144 |
have "{[]} <<| []" |
|
145 |
by simp |
|
146 |
then obtain x :: "'a list" where "{[]} <<| x" .. |
|
147 |
then show ?case |
|
148 |
by auto |
|
39966 | 149 |
next |
68780 | 150 |
case (Cons A B) then show ?case |
151 |
by (auto intro: is_lub_Cons cpo_lubI) |
|
39966 | 152 |
qed |
39143 | 153 |
qed |
154 |
||
62175 | 155 |
subsection \<open>Continuity of list operations\<close> |
39143 | 156 |
|
157 |
lemma cont2cont_Cons [simp, cont2cont]: |
|
158 |
assumes f: "cont (\<lambda>x. f x)" |
|
159 |
assumes g: "cont (\<lambda>x. g x)" |
|
160 |
shows "cont (\<lambda>x. f x # g x)" |
|
161 |
apply (rule contI) |
|
162 |
apply (rule is_lub_Cons) |
|
163 |
apply (erule contE [OF f]) |
|
164 |
apply (erule contE [OF g]) |
|
165 |
done |
|
166 |
||
167 |
lemma lub_Cons: |
|
168 |
fixes A :: "nat \<Rightarrow> 'a::cpo" |
|
169 |
assumes A: "chain A" and B: "chain B" |
|
170 |
shows "(\<Squnion>i. A i # B i) = (\<Squnion>i. A i) # (\<Squnion>i. B i)" |
|
40771 | 171 |
by (intro lub_eqI is_lub_Cons cpo_lubI A B) |
39143 | 172 |
|
55413
a8e96847523c
adapted theories to '{case,rec}_{list,option}' names
blanchet
parents:
54863
diff
changeset
|
173 |
lemma cont2cont_case_list: |
39143 | 174 |
assumes f: "cont (\<lambda>x. f x)" |
175 |
assumes g: "cont (\<lambda>x. g x)" |
|
176 |
assumes h1: "\<And>y ys. cont (\<lambda>x. h x y ys)" |
|
177 |
assumes h2: "\<And>x ys. cont (\<lambda>y. h x y ys)" |
|
178 |
assumes h3: "\<And>x y. cont (\<lambda>ys. h x y ys)" |
|
179 |
shows "cont (\<lambda>x. case f x of [] \<Rightarrow> g x | y # ys \<Rightarrow> h x y ys)" |
|
180 |
apply (rule cont_apply [OF f]) |
|
181 |
apply (rule contI) |
|
182 |
apply (erule list_chain_cases) |
|
40771 | 183 |
apply (simp add: is_lub_const) |
39143 | 184 |
apply (simp add: lub_Cons) |
185 |
apply (simp add: cont2contlubE [OF h2]) |
|
186 |
apply (simp add: cont2contlubE [OF h3]) |
|
187 |
apply (simp add: diag_lub ch2ch_cont [OF h2] ch2ch_cont [OF h3]) |
|
188 |
apply (rule cpo_lubI, rule chainI, rule below_trans) |
|
189 |
apply (erule cont2monofunE [OF h2 chainE]) |
|
190 |
apply (erule cont2monofunE [OF h3 chainE]) |
|
191 |
apply (case_tac y, simp_all add: g h1) |
|
192 |
done |
|
193 |
||
55413
a8e96847523c
adapted theories to '{case,rec}_{list,option}' names
blanchet
parents:
54863
diff
changeset
|
194 |
lemma cont2cont_case_list' [simp, cont2cont]: |
39143 | 195 |
assumes f: "cont (\<lambda>x. f x)" |
196 |
assumes g: "cont (\<lambda>x. g x)" |
|
197 |
assumes h: "cont (\<lambda>p. h (fst p) (fst (snd p)) (snd (snd p)))" |
|
198 |
shows "cont (\<lambda>x. case f x of [] \<Rightarrow> g x | y # ys \<Rightarrow> h x y ys)" |
|
55413
a8e96847523c
adapted theories to '{case,rec}_{list,option}' names
blanchet
parents:
54863
diff
changeset
|
199 |
using assms by (simp add: cont2cont_case_list prod_cont_iff) |
39143 | 200 |
|
62175 | 201 |
text \<open>The simple version (due to Joachim Breitner) is needed if the |
202 |
element type of the list is not a cpo.\<close> |
|
39143 | 203 |
|
55413
a8e96847523c
adapted theories to '{case,rec}_{list,option}' names
blanchet
parents:
54863
diff
changeset
|
204 |
lemma cont2cont_case_list_simple [simp, cont2cont]: |
39143 | 205 |
assumes "cont (\<lambda>x. f1 x)" |
206 |
assumes "\<And>y ys. cont (\<lambda>x. f2 x y ys)" |
|
207 |
shows "cont (\<lambda>x. case l of [] \<Rightarrow> f1 x | y # ys \<Rightarrow> f2 x y ys)" |
|
208 |
using assms by (cases l) auto |
|
209 |
||
62175 | 210 |
text \<open>Lemma for proving continuity of recursive list functions:\<close> |
39966 | 211 |
|
212 |
lemma list_contI: |
|
213 |
fixes f :: "'a::cpo list \<Rightarrow> 'b::cpo" |
|
214 |
assumes f: "\<And>x xs. f (x # xs) = g x xs (f xs)" |
|
215 |
assumes g1: "\<And>xs y. cont (\<lambda>x. g x xs y)" |
|
216 |
assumes g2: "\<And>x y. cont (\<lambda>xs. g x xs y)" |
|
217 |
assumes g3: "\<And>x xs. cont (\<lambda>y. g x xs y)" |
|
218 |
shows "cont f" |
|
219 |
proof (rule contI2) |
|
220 |
obtain h where h: "\<And>x xs y. g x xs y = h\<cdot>x\<cdot>xs\<cdot>y" |
|
221 |
proof |
|
222 |
fix x xs y show "g x xs y = (\<Lambda> x xs y. g x xs y)\<cdot>x\<cdot>xs\<cdot>y" |
|
223 |
by (simp add: cont2cont_LAM g1 g2 g3) |
|
224 |
qed |
|
225 |
show mono: "monofun f" |
|
226 |
apply (rule monofunI) |
|
227 |
apply (erule list_below_induct) |
|
228 |
apply simp |
|
229 |
apply (simp add: f h monofun_cfun) |
|
230 |
done |
|
231 |
fix Y :: "nat \<Rightarrow> 'a list" |
|
232 |
assume "chain Y" thus "f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. f (Y i))" |
|
233 |
apply (induct rule: list_chain_induct) |
|
234 |
apply simp |
|
235 |
apply (simp add: lub_Cons f h) |
|
41027 | 236 |
apply (simp add: lub_APP ch2ch_monofun [OF mono]) |
39966 | 237 |
apply (simp add: monofun_cfun) |
238 |
done |
|
239 |
qed |
|
240 |
||
62175 | 241 |
text \<open>Continuity rule for map\<close> |
40831 | 242 |
|
243 |
lemma cont2cont_map [simp, cont2cont]: |
|
244 |
assumes f: "cont (\<lambda>(x, y). f x y)" |
|
245 |
assumes g: "cont (\<lambda>x. g x)" |
|
246 |
shows "cont (\<lambda>x. map (\<lambda>y. f x y) (g x))" |
|
247 |
using f |
|
248 |
apply (simp add: prod_cont_iff) |
|
249 |
apply (rule cont_apply [OF g]) |
|
55465 | 250 |
apply (rule list_contI, rule list.map, simp, simp, simp) |
40831 | 251 |
apply (induct_tac y, simp, simp) |
252 |
done |
|
253 |
||
62175 | 254 |
text \<open>There are probably lots of other list operations that also |
39143 | 255 |
deserve to have continuity lemmas. I'll add more as they are |
62175 | 256 |
needed.\<close> |
39143 | 257 |
|
62175 | 258 |
subsection \<open>Lists are a discrete cpo\<close> |
40808 | 259 |
|
260 |
instance list :: (discrete_cpo) discrete_cpo |
|
261 |
proof |
|
262 |
fix xs ys :: "'a list" |
|
263 |
show "xs \<sqsubseteq> ys \<longleftrightarrow> xs = ys" |
|
264 |
by (induct xs arbitrary: ys, case_tac [!] ys, simp_all) |
|
265 |
qed |
|
266 |
||
62175 | 267 |
subsection \<open>Compactness and chain-finiteness\<close> |
40808 | 268 |
|
269 |
lemma compact_Nil [simp]: "compact []" |
|
270 |
apply (rule compactI2) |
|
271 |
apply (erule list_chain_cases) |
|
272 |
apply simp |
|
273 |
apply (simp add: lub_Cons) |
|
274 |
done |
|
275 |
||
276 |
lemma compact_Cons: "\<lbrakk>compact x; compact xs\<rbrakk> \<Longrightarrow> compact (x # xs)" |
|
277 |
apply (rule compactI2) |
|
278 |
apply (erule list_chain_cases) |
|
279 |
apply (auto simp add: lub_Cons dest!: compactD2) |
|
280 |
apply (rename_tac i j, rule_tac x="max i j" in exI) |
|
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
42151
diff
changeset
|
281 |
apply (drule (1) below_trans [OF _ chain_mono [OF _ max.cobounded1]]) |
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
42151
diff
changeset
|
282 |
apply (drule (1) below_trans [OF _ chain_mono [OF _ max.cobounded2]]) |
40808 | 283 |
apply (erule (1) conjI) |
284 |
done |
|
285 |
||
286 |
lemma compact_Cons_iff [simp]: |
|
287 |
"compact (x # xs) \<longleftrightarrow> compact x \<and> compact xs" |
|
288 |
apply (safe intro!: compact_Cons) |
|
289 |
apply (simp only: compact_def) |
|
290 |
apply (subgoal_tac "cont (\<lambda>x. x # xs)") |
|
291 |
apply (drule (1) adm_subst, simp, simp) |
|
292 |
apply (simp only: compact_def) |
|
293 |
apply (subgoal_tac "cont (\<lambda>xs. x # xs)") |
|
294 |
apply (drule (1) adm_subst, simp, simp) |
|
295 |
done |
|
296 |
||
297 |
instance list :: (chfin) chfin |
|
298 |
proof |
|
299 |
fix Y :: "nat \<Rightarrow> 'a list" assume "chain Y" |
|
300 |
moreover have "\<And>(xs::'a list). compact xs" |
|
301 |
by (induct_tac xs, simp_all) |
|
302 |
ultimately show "\<exists>i. max_in_chain i Y" |
|
303 |
by (rule compact_imp_max_in_chain) |
|
304 |
qed |
|
305 |
||
62175 | 306 |
subsection \<open>Using lists with fixrec\<close> |
39212 | 307 |
|
308 |
definition |
|
309 |
match_Nil :: "'a::cpo list \<rightarrow> 'b match \<rightarrow> 'b match" |
|
310 |
where |
|
311 |
"match_Nil = (\<Lambda> xs k. case xs of [] \<Rightarrow> k | y # ys \<Rightarrow> Fixrec.fail)" |
|
312 |
||
313 |
definition |
|
314 |
match_Cons :: "'a::cpo list \<rightarrow> ('a \<rightarrow> 'a list \<rightarrow> 'b match) \<rightarrow> 'b match" |
|
315 |
where |
|
316 |
"match_Cons = (\<Lambda> xs k. case xs of [] \<Rightarrow> Fixrec.fail | y # ys \<Rightarrow> k\<cdot>y\<cdot>ys)" |
|
317 |
||
318 |
lemma match_Nil_simps [simp]: |
|
319 |
"match_Nil\<cdot>[]\<cdot>k = k" |
|
320 |
"match_Nil\<cdot>(x # xs)\<cdot>k = Fixrec.fail" |
|
321 |
unfolding match_Nil_def by simp_all |
|
322 |
||
323 |
lemma match_Cons_simps [simp]: |
|
324 |
"match_Cons\<cdot>[]\<cdot>k = Fixrec.fail" |
|
325 |
"match_Cons\<cdot>(x # xs)\<cdot>k = k\<cdot>x\<cdot>xs" |
|
326 |
unfolding match_Cons_def by simp_all |
|
327 |
||
62175 | 328 |
setup \<open> |
39212 | 329 |
Fixrec.add_matchers |
69597 | 330 |
[ (\<^const_name>\<open>Nil\<close>, \<^const_name>\<open>match_Nil\<close>), |
331 |
(\<^const_name>\<open>Cons\<close>, \<^const_name>\<open>match_Cons\<close>) ] |
|
62175 | 332 |
\<close> |
39212 | 333 |
|
39143 | 334 |
end |