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