| author | wenzelm | 
| Tue, 07 Mar 2023 11:13:36 +0100 | |
| changeset 77555 | d45a01c41fe2 | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/Product_Cpo.thy | 
| 29531 | 2 | Author: Franz Regensburger | 
| 3 | *) | |
| 4 | ||
| 62175 | 5 | section \<open>The cpo of cartesian products\<close> | 
| 29531 | 6 | |
| 7 | theory Product_Cpo | |
| 67312 | 8 | imports Adm | 
| 29531 | 9 | begin | 
| 10 | ||
| 36452 | 11 | default_sort cpo | 
| 29531 | 12 | |
| 67312 | 13 | |
| 62175 | 14 | subsection \<open>Unit type is a pcpo\<close> | 
| 29531 | 15 | |
| 40090 | 16 | instantiation unit :: discrete_cpo | 
| 29531 | 17 | begin | 
| 18 | ||
| 67312 | 19 | definition below_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<longleftrightarrow> True" | 
| 29531 | 20 | |
| 67312 | 21 | instance | 
| 22 | by standard simp | |
| 29531 | 23 | |
| 40090 | 24 | end | 
| 29531 | 25 | |
| 26 | instance unit :: pcpo | |
| 67312 | 27 | by standard simp | 
| 29531 | 28 | |
| 29 | ||
| 62175 | 30 | subsection \<open>Product type is a partial order\<close> | 
| 29531 | 31 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37079diff
changeset | 32 | instantiation prod :: (below, below) below | 
| 29531 | 33 | begin | 
| 34 | ||
| 67399 | 35 | definition below_prod_def: "(\<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)" | 
| 29531 | 36 | |
| 37 | instance .. | |
| 67312 | 38 | |
| 29531 | 39 | end | 
| 40 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37079diff
changeset | 41 | instance prod :: (po, po) po | 
| 29531 | 42 | proof | 
| 43 | fix x :: "'a \<times> 'b" | |
| 44 | show "x \<sqsubseteq> x" | |
| 67312 | 45 | by (simp add: below_prod_def) | 
| 29531 | 46 | next | 
| 47 | fix x y :: "'a \<times> 'b" | |
| 67312 | 48 | assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" | 
| 49 | then show "x = y" | |
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
42151diff
changeset | 50 | unfolding below_prod_def prod_eq_iff | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31041diff
changeset | 51 | by (fast intro: below_antisym) | 
| 29531 | 52 | next | 
| 53 | fix x y z :: "'a \<times> 'b" | |
| 67312 | 54 | assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" | 
| 55 | then show "x \<sqsubseteq> z" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31041diff
changeset | 56 | unfolding below_prod_def | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31041diff
changeset | 57 | by (fast intro: below_trans) | 
| 29531 | 58 | qed | 
| 59 | ||
| 67312 | 60 | |
| 62175 | 61 | subsection \<open>Monotonicity of \emph{Pair}, \emph{fst}, \emph{snd}\<close>
 | 
| 29531 | 62 | |
| 67312 | 63 | lemma prod_belowI: "fst p \<sqsubseteq> fst q \<Longrightarrow> snd p \<sqsubseteq> snd q \<Longrightarrow> p \<sqsubseteq> q" | 
| 64 | by (simp add: below_prod_def) | |
| 29531 | 65 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31041diff
changeset | 66 | lemma Pair_below_iff [simp]: "(a, b) \<sqsubseteq> (c, d) \<longleftrightarrow> a \<sqsubseteq> c \<and> b \<sqsubseteq> d" | 
| 67312 | 67 | by (simp add: below_prod_def) | 
| 29531 | 68 | |
| 62175 | 69 | text \<open>Pair \<open>(_,_)\<close> is monotone in both arguments\<close> | 
| 29531 | 70 | |
| 71 | lemma monofun_pair1: "monofun (\<lambda>x. (x, y))" | |
| 67312 | 72 | by (simp add: monofun_def) | 
| 29531 | 73 | |
| 74 | lemma monofun_pair2: "monofun (\<lambda>y. (x, y))" | |
| 67312 | 75 | by (simp add: monofun_def) | 
| 29531 | 76 | |
| 67312 | 77 | lemma monofun_pair: "x1 \<sqsubseteq> x2 \<Longrightarrow> y1 \<sqsubseteq> y2 \<Longrightarrow> (x1, y1) \<sqsubseteq> (x2, y2)" | 
| 78 | by simp | |
| 29531 | 79 | |
| 67312 | 80 | lemma ch2ch_Pair [simp]: "chain X \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. (X i, Y i))" | 
| 81 | by (rule chainI, simp add: chainE) | |
| 31112 | 82 | |
| 69597 | 83 | text \<open>\<^term>\<open>fst\<close> and \<^term>\<open>snd\<close> are monotone\<close> | 
| 29531 | 84 | |
| 35919 | 85 | lemma fst_monofun: "x \<sqsubseteq> y \<Longrightarrow> fst x \<sqsubseteq> fst y" | 
| 67312 | 86 | by (simp add: below_prod_def) | 
| 35919 | 87 | |
| 88 | lemma snd_monofun: "x \<sqsubseteq> y \<Longrightarrow> snd x \<sqsubseteq> snd y" | |
| 67312 | 89 | by (simp add: below_prod_def) | 
| 35919 | 90 | |
| 29531 | 91 | lemma monofun_fst: "monofun fst" | 
| 67312 | 92 | by (simp add: monofun_def below_prod_def) | 
| 29531 | 93 | |
| 94 | lemma monofun_snd: "monofun snd" | |
| 67312 | 95 | by (simp add: monofun_def below_prod_def) | 
| 29531 | 96 | |
| 31112 | 97 | lemmas ch2ch_fst [simp] = ch2ch_monofun [OF monofun_fst] | 
| 98 | ||
| 99 | lemmas ch2ch_snd [simp] = ch2ch_monofun [OF monofun_snd] | |
| 100 | ||
| 101 | lemma prod_chain_cases: | |
| 67312 | 102 | assumes chain: "chain Y" | 
| 31112 | 103 | obtains A B | 
| 104 | where "chain A" and "chain B" and "Y = (\<lambda>i. (A i, B i))" | |
| 105 | proof | |
| 67312 | 106 | from chain show "chain (\<lambda>i. fst (Y i))" | 
| 107 | by (rule ch2ch_fst) | |
| 108 | from chain show "chain (\<lambda>i. snd (Y i))" | |
| 109 | by (rule ch2ch_snd) | |
| 110 | show "Y = (\<lambda>i. (fst (Y i), snd (Y i)))" | |
| 111 | by simp | |
| 31112 | 112 | qed | 
| 113 | ||
| 67312 | 114 | |
| 62175 | 115 | subsection \<open>Product type is a cpo\<close> | 
| 29531 | 116 | |
| 67312 | 117 | lemma is_lub_Pair: "range A <<| x \<Longrightarrow> range B <<| y \<Longrightarrow> range (\<lambda>i. (A i, B i)) <<| (x, y)" | 
| 118 | by (simp add: is_lub_def is_ub_def below_prod_def) | |
| 29531 | 119 | |
| 67312 | 120 | lemma lub_Pair: "chain A \<Longrightarrow> chain B \<Longrightarrow> (\<Squnion>i. (A i, B i)) = (\<Squnion>i. A i, \<Squnion>i. B i)" | 
| 121 | for A :: "nat \<Rightarrow> 'a::cpo" and B :: "nat \<Rightarrow> 'b::cpo" | |
| 122 | by (fast intro: lub_eqI is_lub_Pair elim: thelubE) | |
| 31112 | 123 | |
| 40771 | 124 | lemma is_lub_prod: | 
| 29531 | 125 |   fixes S :: "nat \<Rightarrow> ('a::cpo \<times> 'b::cpo)"
 | 
| 67312 | 126 | assumes "chain S" | 
| 29531 | 127 | shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" | 
| 67312 | 128 | using assms by (auto elim: prod_chain_cases simp: is_lub_Pair cpo_lubI) | 
| 29531 | 129 | |
| 67312 | 130 | lemma lub_prod: "chain S \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" | 
| 131 | for S :: "nat \<Rightarrow> 'a::cpo \<times> 'b::cpo" | |
| 132 | by (rule is_lub_prod [THEN lub_eqI]) | |
| 29531 | 133 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37079diff
changeset | 134 | instance prod :: (cpo, cpo) cpo | 
| 29531 | 135 | proof | 
| 136 |   fix S :: "nat \<Rightarrow> ('a \<times> 'b)"
 | |
| 137 | assume "chain S" | |
| 67312 | 138 | then have "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" | 
| 40771 | 139 | by (rule is_lub_prod) | 
| 67312 | 140 | then show "\<exists>x. range S <<| x" .. | 
| 29531 | 141 | qed | 
| 142 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37079diff
changeset | 143 | instance prod :: (discrete_cpo, discrete_cpo) discrete_cpo | 
| 29531 | 144 | proof | 
| 145 | fix x y :: "'a \<times> 'b" | |
| 146 | show "x \<sqsubseteq> y \<longleftrightarrow> x = y" | |
| 67312 | 147 | by (simp add: below_prod_def prod_eq_iff) | 
| 29531 | 148 | qed | 
| 149 | ||
| 67312 | 150 | |
| 62175 | 151 | subsection \<open>Product type is pointed\<close> | 
| 29531 | 152 | |
| 40771 | 153 | lemma minimal_prod: "(\<bottom>, \<bottom>) \<sqsubseteq> p" | 
| 67312 | 154 | by (simp add: below_prod_def) | 
| 29531 | 155 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37079diff
changeset | 156 | instance prod :: (pcpo, pcpo) pcpo | 
| 67312 | 157 | by intro_classes (fast intro: minimal_prod) | 
| 29531 | 158 | |
| 40771 | 159 | lemma inst_prod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)" | 
| 67312 | 160 | by (rule minimal_prod [THEN bottomI, symmetric]) | 
| 29531 | 161 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40090diff
changeset | 162 | lemma Pair_bottom_iff [simp]: "(x, y) = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>" | 
| 67312 | 163 | by (simp add: inst_prod_pcpo) | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 164 | |
| 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 165 | lemma fst_strict [simp]: "fst \<bottom> = \<bottom>" | 
| 67312 | 166 | unfolding inst_prod_pcpo by (rule fst_conv) | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 167 | |
| 33506 | 168 | lemma snd_strict [simp]: "snd \<bottom> = \<bottom>" | 
| 67312 | 169 | unfolding inst_prod_pcpo by (rule snd_conv) | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 170 | |
| 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 171 | lemma Pair_strict [simp]: "(\<bottom>, \<bottom>) = \<bottom>" | 
| 67312 | 172 | by simp | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 173 | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
58880diff
changeset | 174 | lemma split_strict [simp]: "case_prod f \<bottom> = f \<bottom> \<bottom>" | 
| 67312 | 175 | by (simp add: split_def) | 
| 176 | ||
| 29531 | 177 | |
| 62175 | 178 | subsection \<open>Continuity of \emph{Pair}, \emph{fst}, \emph{snd}\<close>
 | 
| 29531 | 179 | |
| 180 | lemma cont_pair1: "cont (\<lambda>x. (x, y))" | |
| 67312 | 181 | apply (rule contI) | 
| 182 | apply (rule is_lub_Pair) | |
| 183 | apply (erule cpo_lubI) | |
| 184 | apply (rule is_lub_const) | |
| 185 | done | |
| 29531 | 186 | |
| 187 | lemma cont_pair2: "cont (\<lambda>y. (x, y))" | |
| 67312 | 188 | apply (rule contI) | 
| 189 | apply (rule is_lub_Pair) | |
| 190 | apply (rule is_lub_const) | |
| 191 | apply (erule cpo_lubI) | |
| 192 | done | |
| 29531 | 193 | |
| 35914 | 194 | lemma cont_fst: "cont fst" | 
| 67312 | 195 | apply (rule contI) | 
| 196 | apply (simp add: lub_prod) | |
| 197 | apply (erule cpo_lubI [OF ch2ch_fst]) | |
| 198 | done | |
| 29531 | 199 | |
| 200 | lemma cont_snd: "cont snd" | |
| 67312 | 201 | apply (rule contI) | 
| 202 | apply (simp add: lub_prod) | |
| 203 | apply (erule cpo_lubI [OF ch2ch_snd]) | |
| 204 | done | |
| 29531 | 205 | |
| 37079 
0cd15d8c90a0
remove cont2cont simproc; instead declare cont2cont rules as simp rules
 huffman parents: 
36452diff
changeset | 206 | lemma cont2cont_Pair [simp, cont2cont]: | 
| 29531 | 207 | assumes f: "cont (\<lambda>x. f x)" | 
| 208 | assumes g: "cont (\<lambda>x. g x)" | |
| 209 | shows "cont (\<lambda>x. (f x, g x))" | |
| 67312 | 210 | apply (rule cont_apply [OF f cont_pair1]) | 
| 211 | apply (rule cont_apply [OF g cont_pair2]) | |
| 212 | apply (rule cont_const) | |
| 213 | done | |
| 29531 | 214 | |
| 37079 
0cd15d8c90a0
remove cont2cont simproc; instead declare cont2cont rules as simp rules
 huffman parents: 
36452diff
changeset | 215 | lemmas cont2cont_fst [simp, cont2cont] = cont_compose [OF cont_fst] | 
| 31041 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 216 | |
| 37079 
0cd15d8c90a0
remove cont2cont simproc; instead declare cont2cont rules as simp rules
 huffman parents: 
36452diff
changeset | 217 | lemmas cont2cont_snd [simp, cont2cont] = cont_compose [OF cont_snd] | 
| 31041 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 218 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
44066diff
changeset | 219 | lemma cont2cont_case_prod: | 
| 31041 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 220 | assumes f1: "\<And>a b. cont (\<lambda>x. f x a b)" | 
| 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 221 | assumes f2: "\<And>x b. cont (\<lambda>a. f x a b)" | 
| 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 222 | assumes f3: "\<And>x a. cont (\<lambda>b. f x a b)" | 
| 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 223 | assumes g: "cont (\<lambda>x. g x)" | 
| 39144 | 224 | shows "cont (\<lambda>x. case g x of (a, b) \<Rightarrow> f x a b)" | 
| 67312 | 225 | unfolding split_def | 
| 226 | apply (rule cont_apply [OF g]) | |
| 227 | apply (rule cont_apply [OF cont_fst f2]) | |
| 228 | apply (rule cont_apply [OF cont_snd f3]) | |
| 229 | apply (rule cont_const) | |
| 230 | apply (rule f1) | |
| 231 | done | |
| 31041 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 232 | |
| 39808 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 233 | lemma prod_contI: | 
| 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 234 | assumes f1: "\<And>y. cont (\<lambda>x. f (x, y))" | 
| 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 235 | assumes f2: "\<And>x. cont (\<lambda>y. f (x, y))" | 
| 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 236 | shows "cont f" | 
| 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 237 | proof - | 
| 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 238 | have "cont (\<lambda>(x, y). f (x, y))" | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
44066diff
changeset | 239 | by (intro cont2cont_case_prod f1 f2 cont2cont) | 
| 67312 | 240 | then show "cont f" | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
58880diff
changeset | 241 | by (simp only: case_prod_eta) | 
| 39808 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 242 | qed | 
| 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 243 | |
| 67312 | 244 | lemma prod_cont_iff: "cont f \<longleftrightarrow> (\<forall>y. cont (\<lambda>x. f (x, y))) \<and> (\<forall>x. cont (\<lambda>y. f (x, y)))" | 
| 245 | apply safe | |
| 246 | apply (erule cont_compose [OF _ cont_pair1]) | |
| 247 | apply (erule cont_compose [OF _ cont_pair2]) | |
| 248 | apply (simp only: prod_contI) | |
| 249 | done | |
| 39808 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 huffman parents: 
39144diff
changeset | 250 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
44066diff
changeset | 251 | lemma cont2cont_case_prod' [simp, cont2cont]: | 
| 31041 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 252 | assumes f: "cont (\<lambda>p. f (fst p) (fst (snd p)) (snd (snd p)))" | 
| 
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
 huffman parents: 
29535diff
changeset | 253 | assumes g: "cont (\<lambda>x. g x)" | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
44066diff
changeset | 254 | shows "cont (\<lambda>x. case_prod (f x) (g x))" | 
| 67312 | 255 | using assms by (simp add: cont2cont_case_prod prod_cont_iff) | 
| 29531 | 256 | |
| 62175 | 257 | text \<open>The simple version (due to Joachim Breitner) is needed if | 
| 258 | either element type of the pair is not a cpo.\<close> | |
| 39144 | 259 | |
| 260 | lemma cont2cont_split_simple [simp, cont2cont]: | |
| 67312 | 261 | assumes "\<And>a b. cont (\<lambda>x. f x a b)" | 
| 262 | shows "cont (\<lambda>x. case p of (a, b) \<Rightarrow> f x a b)" | |
| 263 | using assms by (cases p) auto | |
| 39144 | 264 | |
| 62175 | 265 | text \<open>Admissibility of predicates on product types.\<close> | 
| 40619 | 266 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
44066diff
changeset | 267 | lemma adm_case_prod [simp]: | 
| 40619 | 268 | assumes "adm (\<lambda>x. P x (fst (f x)) (snd (f x)))" | 
| 269 | shows "adm (\<lambda>x. case f x of (a, b) \<Rightarrow> P x a b)" | |
| 67312 | 270 | unfolding case_prod_beta using assms . | 
| 271 | ||
| 40619 | 272 | |
| 62175 | 273 | subsection \<open>Compactness and chain-finiteness\<close> | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 274 | |
| 67312 | 275 | lemma fst_below_iff: "fst x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (y, snd x)" | 
| 276 | for x :: "'a \<times> 'b" | |
| 277 | by (simp add: below_prod_def) | |
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 278 | |
| 67312 | 279 | lemma snd_below_iff: "snd x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (fst x, y)" | 
| 280 | for x :: "'a \<times> 'b" | |
| 281 | by (simp add: below_prod_def) | |
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 282 | |
| 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 283 | lemma compact_fst: "compact x \<Longrightarrow> compact (fst x)" | 
| 67312 | 284 | by (rule compactI) (simp add: fst_below_iff) | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 285 | |
| 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 286 | lemma compact_snd: "compact x \<Longrightarrow> compact (snd x)" | 
| 67312 | 287 | by (rule compactI) (simp add: snd_below_iff) | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 288 | |
| 67312 | 289 | lemma compact_Pair: "compact x \<Longrightarrow> compact y \<Longrightarrow> compact (x, y)" | 
| 290 | by (rule compactI) (simp add: below_prod_def) | |
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 291 | |
| 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 292 | lemma compact_Pair_iff [simp]: "compact (x, y) \<longleftrightarrow> compact x \<and> compact y" | 
| 67312 | 293 | apply (safe intro!: compact_Pair) | 
| 294 | apply (drule compact_fst, simp) | |
| 295 | apply (drule compact_snd, simp) | |
| 296 | done | |
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 297 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37079diff
changeset | 298 | instance prod :: (chfin, chfin) chfin | 
| 67312 | 299 | apply intro_classes | 
| 300 | apply (erule compact_imp_max_in_chain) | |
| 301 | apply (case_tac "\<Squnion>i. Y i", simp) | |
| 302 | done | |
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29533diff
changeset | 303 | |
| 29531 | 304 | end |