| author | wenzelm | 
| Sat, 23 Dec 2023 16:12:53 +0100 | |
| changeset 79339 | 8eb7e325f40d | 
| parent 76952 | f552cf789a8d | 
| child 80760 | be8c0e039a5e | 
| permissions | -rw-r--r-- | 
| 55056 | 1  | 
(* Title: HOL/BNF_Cardinal_Arithmetic.thy  | 
| 54474 | 2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
3  | 
Copyright 2012  | 
|
4  | 
||
| 55059 | 5  | 
Cardinal arithmetic as needed by bounded natural functors.  | 
| 54474 | 6  | 
*)  | 
7  | 
||
| 60758 | 8  | 
section \<open>Cardinal Arithmetic as Needed by Bounded Natural Functors\<close>  | 
| 54474 | 9  | 
|
| 55056 | 10  | 
theory BNF_Cardinal_Arithmetic  | 
| 76951 | 11  | 
imports BNF_Cardinal_Order_Relation  | 
| 54474 | 12  | 
begin  | 
13  | 
||
14  | 
lemma dir_image: "\<lbrakk>\<And>x y. (f x = f y) = (x = y); Card_order r\<rbrakk> \<Longrightarrow> r =o dir_image r f"  | 
|
| 76951 | 15  | 
by (rule dir_image_ordIso) (auto simp add: inj_on_def card_order_on_def)  | 
| 54474 | 16  | 
|
17  | 
lemma card_order_dir_image:  | 
|
18  | 
assumes bij: "bij f" and co: "card_order r"  | 
|
19  | 
shows "card_order (dir_image r f)"  | 
|
20  | 
proof -  | 
|
| 76951 | 21  | 
have "Field (dir_image r f) = UNIV"  | 
22  | 
using assms card_order_on_Card_order[of UNIV r]  | 
|
23  | 
unfolding bij_def dir_image_Field by auto  | 
|
24  | 
moreover from bij have "\<And>x y. (f x = f y) = (x = y)"  | 
|
25  | 
unfolding bij_def inj_on_def by auto  | 
|
| 54474 | 26  | 
with co have "Card_order (dir_image r f)"  | 
| 76951 | 27  | 
using card_order_on_Card_order Card_order_ordIso2[OF _ dir_image] by blast  | 
| 54474 | 28  | 
ultimately show ?thesis by auto  | 
29  | 
qed  | 
|
30  | 
||
31  | 
lemma ordIso_refl: "Card_order r \<Longrightarrow> r =o r"  | 
|
| 76951 | 32  | 
by (rule card_order_on_ordIso)  | 
| 54474 | 33  | 
|
34  | 
lemma ordLeq_refl: "Card_order r \<Longrightarrow> r \<le>o r"  | 
|
| 76951 | 35  | 
by (rule ordIso_imp_ordLeq, rule card_order_on_ordIso)  | 
| 54474 | 36  | 
|
37  | 
lemma card_of_ordIso_subst: "A = B \<Longrightarrow> |A| =o |B|"  | 
|
| 76951 | 38  | 
by (simp only: ordIso_refl card_of_Card_order)  | 
| 54474 | 39  | 
|
40  | 
lemma Field_card_order: "card_order r \<Longrightarrow> Field r = UNIV"  | 
|
| 76951 | 41  | 
using card_order_on_Card_order[of UNIV r] by simp  | 
| 54474 | 42  | 
|
43  | 
||
| 60758 | 44  | 
subsection \<open>Zero\<close>  | 
| 54474 | 45  | 
|
46  | 
definition czero where  | 
|
47  | 
  "czero = card_of {}"
 | 
|
48  | 
||
| 76951 | 49  | 
lemma czero_ordIso: "czero =o czero"  | 
50  | 
using card_of_empty_ordIso by (simp add: czero_def)  | 
|
| 54474 | 51  | 
|
52  | 
lemma card_of_ordIso_czero_iff_empty:  | 
|
53  | 
  "|A| =o (czero :: 'b rel) \<longleftrightarrow> A = ({} :: 'a set)"
 | 
|
| 76951 | 54  | 
unfolding czero_def by (rule iffI[OF card_of_empty2]) (auto simp: card_of_refl card_of_empty_ordIso)  | 
| 54474 | 55  | 
|
56  | 
(* A "not czero" Cardinal predicate *)  | 
|
57  | 
abbreviation Cnotzero where  | 
|
58  | 
"Cnotzero (r :: 'a rel) \<equiv> \<not>(r =o (czero :: 'a rel)) \<and> Card_order r"  | 
|
59  | 
||
60  | 
(*helper*)  | 
|
61  | 
lemma Cnotzero_imp_not_empty: "Cnotzero r \<Longrightarrow> Field r \<noteq> {}"
 | 
|
| 55811 | 62  | 
unfolding Card_order_iff_ordIso_card_of czero_def by force  | 
| 54474 | 63  | 
|
64  | 
lemma czeroI:  | 
|
65  | 
  "\<lbrakk>Card_order r; Field r = {}\<rbrakk> \<Longrightarrow> r =o czero"
 | 
|
| 76951 | 66  | 
using Cnotzero_imp_not_empty ordIso_transitive[OF _ czero_ordIso] by blast  | 
| 54474 | 67  | 
|
68  | 
lemma czeroE:  | 
|
69  | 
  "r =o czero \<Longrightarrow> Field r = {}"
 | 
|
| 76951 | 70  | 
unfolding czero_def  | 
71  | 
by (drule card_of_cong) (simp only: Field_card_of card_of_empty2)  | 
|
| 54474 | 72  | 
|
73  | 
lemma Cnotzero_mono:  | 
|
74  | 
"\<lbrakk>Cnotzero r; Card_order q; r \<le>o q\<rbrakk> \<Longrightarrow> Cnotzero q"  | 
|
| 76951 | 75  | 
by (force intro: czeroI dest: card_of_mono2 card_of_empty3 czeroE)  | 
| 54474 | 76  | 
|
| 60758 | 77  | 
subsection \<open>(In)finite cardinals\<close>  | 
| 54474 | 78  | 
|
79  | 
definition cinfinite where  | 
|
| 76951 | 80  | 
"cinfinite r \<equiv> (\<not> finite (Field r))"  | 
| 54474 | 81  | 
|
82  | 
abbreviation Cinfinite where  | 
|
83  | 
"Cinfinite r \<equiv> cinfinite r \<and> Card_order r"  | 
|
84  | 
||
85  | 
definition cfinite where  | 
|
86  | 
"cfinite r = finite (Field r)"  | 
|
87  | 
||
88  | 
abbreviation Cfinite where  | 
|
89  | 
"Cfinite r \<equiv> cfinite r \<and> Card_order r"  | 
|
90  | 
||
91  | 
lemma Cfinite_ordLess_Cinfinite: "\<lbrakk>Cfinite r; Cinfinite s\<rbrakk> \<Longrightarrow> r <o s"  | 
|
92  | 
unfolding cfinite_def cinfinite_def  | 
|
| 55811 | 93  | 
by (blast intro: finite_ordLess_infinite card_order_on_well_order_on)  | 
| 54474 | 94  | 
|
| 
54581
 
1502a1f707d9
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
 
traytel 
parents: 
54578 
diff
changeset
 | 
95  | 
lemmas natLeq_card_order = natLeq_Card_order[unfolded Field_natLeq]  | 
| 
 
1502a1f707d9
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
 
traytel 
parents: 
54578 
diff
changeset
 | 
96  | 
|
| 
 
1502a1f707d9
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
 
traytel 
parents: 
54578 
diff
changeset
 | 
97  | 
lemma natLeq_cinfinite: "cinfinite natLeq"  | 
| 76951 | 98  | 
unfolding cinfinite_def Field_natLeq by (rule infinite_UNIV_nat)  | 
| 
54581
 
1502a1f707d9
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
 
traytel 
parents: 
54578 
diff
changeset
 | 
99  | 
|
| 75624 | 100  | 
lemma natLeq_Cinfinite: "Cinfinite natLeq"  | 
101  | 
using natLeq_cinfinite natLeq_Card_order by simp  | 
|
102  | 
||
| 54474 | 103  | 
lemma natLeq_ordLeq_cinfinite:  | 
104  | 
assumes inf: "Cinfinite r"  | 
|
105  | 
shows "natLeq \<le>o r"  | 
|
106  | 
proof -  | 
|
| 55811 | 107  | 
from inf have "natLeq \<le>o |Field r|" unfolding cinfinite_def  | 
108  | 
using infinite_iff_natLeq_ordLeq by blast  | 
|
| 54474 | 109  | 
also from inf have "|Field r| =o r" by (simp add: card_of_unique ordIso_symmetric)  | 
110  | 
finally show ?thesis .  | 
|
111  | 
qed  | 
|
112  | 
||
113  | 
lemma cinfinite_not_czero: "cinfinite r \<Longrightarrow> \<not> (r =o (czero :: 'a rel))"  | 
|
| 76951 | 114  | 
  unfolding cinfinite_def by (cases "Field r = {}") (auto dest: czeroE)
 | 
| 54474 | 115  | 
|
116  | 
lemma Cinfinite_Cnotzero: "Cinfinite r \<Longrightarrow> Cnotzero r"  | 
|
| 76951 | 117  | 
using cinfinite_not_czero by auto  | 
| 54474 | 118  | 
|
119  | 
lemma Cinfinite_cong: "\<lbrakk>r1 =o r2; Cinfinite r1\<rbrakk> \<Longrightarrow> Cinfinite r2"  | 
|
| 76951 | 120  | 
using Card_order_ordIso2[of r1 r2] unfolding cinfinite_def ordIso_iff_ordLeq  | 
121  | 
by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2])  | 
|
| 54474 | 122  | 
|
123  | 
lemma cinfinite_mono: "\<lbrakk>r1 \<le>o r2; cinfinite r1\<rbrakk> \<Longrightarrow> cinfinite r2"  | 
|
| 76951 | 124  | 
unfolding cinfinite_def by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2])  | 
| 54474 | 125  | 
|
| 75624 | 126  | 
lemma regularCard_ordIso:  | 
| 76951 | 127  | 
assumes "k =o k'" and "Cinfinite k" and "regularCard k"  | 
128  | 
shows "regularCard k'"  | 
|
| 75624 | 129  | 
proof-  | 
130  | 
have "stable k" using assms cinfinite_def regularCard_stable by blast  | 
|
131  | 
hence "stable k'" using assms stable_ordIso1 ordIso_symmetric by blast  | 
|
| 76951 | 132  | 
thus ?thesis using assms cinfinite_def stable_regularCard Cinfinite_cong by blast  | 
| 75624 | 133  | 
qed  | 
134  | 
||
135  | 
corollary card_of_UNION_ordLess_infinite_Field_regularCard:  | 
|
| 76951 | 136  | 
assumes "regularCard r" and "Cinfinite r" and "|I| <o r" and "\<forall>i \<in> I. |A i| <o r"  | 
137  | 
shows "|\<Union>i \<in> I. A i| <o r"  | 
|
| 75624 | 138  | 
using card_of_UNION_ordLess_infinite_Field regularCard_stable assms cinfinite_def by blast  | 
| 54474 | 139  | 
|
| 60758 | 140  | 
subsection \<open>Binary sum\<close>  | 
| 54474 | 141  | 
|
| 76951 | 142  | 
definition csum (infixr "+c" 65)  | 
143  | 
where "r1 +c r2 \<equiv> |Field r1 <+> Field r2|"  | 
|
| 54474 | 144  | 
|
145  | 
lemma Field_csum: "Field (r +c s) = Inl ` Field r \<union> Inr ` Field s"  | 
|
146  | 
unfolding csum_def Field_card_of by auto  | 
|
147  | 
||
| 76951 | 148  | 
lemma Card_order_csum: "Card_order (r1 +c r2)"  | 
149  | 
unfolding csum_def by (simp add: card_of_Card_order)  | 
|
| 54474 | 150  | 
|
| 76951 | 151  | 
lemma csum_Cnotzero1: "Cnotzero r1 \<Longrightarrow> Cnotzero (r1 +c r2)"  | 
152  | 
using Cnotzero_imp_not_empty  | 
|
153  | 
by (auto intro: card_of_Card_order simp: csum_def card_of_ordIso_czero_iff_empty)  | 
|
| 54474 | 154  | 
|
155  | 
lemma card_order_csum:  | 
|
156  | 
assumes "card_order r1" "card_order r2"  | 
|
157  | 
shows "card_order (r1 +c r2)"  | 
|
158  | 
proof -  | 
|
159  | 
have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto  | 
|
160  | 
thus ?thesis unfolding csum_def by (auto simp: card_of_card_order_on)  | 
|
161  | 
qed  | 
|
162  | 
||
163  | 
lemma cinfinite_csum:  | 
|
164  | 
"cinfinite r1 \<or> cinfinite r2 \<Longrightarrow> cinfinite (r1 +c r2)"  | 
|
| 76951 | 165  | 
unfolding cinfinite_def csum_def by (auto simp: Field_card_of)  | 
| 54474 | 166  | 
|
| 
54480
 
57e781b711b5
no need for 3-way split with GFP for a handful of theorems
 
blanchet 
parents: 
54474 
diff
changeset
 | 
167  | 
lemma Cinfinite_csum:  | 
| 
 
57e781b711b5
no need for 3-way split with GFP for a handful of theorems
 
blanchet 
parents: 
54474 
diff
changeset
 | 
168  | 
"Cinfinite r1 \<or> Cinfinite r2 \<Longrightarrow> Cinfinite (r1 +c r2)"  | 
| 76951 | 169  | 
using card_of_Card_order  | 
170  | 
by (auto simp: cinfinite_def csum_def Field_card_of)  | 
|
| 
54480
 
57e781b711b5
no need for 3-way split with GFP for a handful of theorems
 
blanchet 
parents: 
54474 
diff
changeset
 | 
171  | 
|
| 76951 | 172  | 
lemma Cinfinite_csum1: "Cinfinite r1 \<Longrightarrow> Cinfinite (r1 +c r2)"  | 
173  | 
by (blast intro!: Cinfinite_csum elim: )  | 
|
| 
54480
 
57e781b711b5
no need for 3-way split with GFP for a handful of theorems
 
blanchet 
parents: 
54474 
diff
changeset
 | 
174  | 
|
| 76952 | 175  | 
lemma Cinfinite_csum_weak:  | 
176  | 
"\<lbrakk>Cinfinite r1; Cinfinite r2\<rbrakk> \<Longrightarrow> Cinfinite (r1 +c r2)"  | 
|
177  | 
by (erule Cinfinite_csum1)  | 
|
178  | 
||
| 54474 | 179  | 
lemma csum_cong: "\<lbrakk>p1 =o r1; p2 =o r2\<rbrakk> \<Longrightarrow> p1 +c p2 =o r1 +c r2"  | 
| 76951 | 180  | 
by (simp only: csum_def ordIso_Plus_cong)  | 
| 54474 | 181  | 
|
182  | 
lemma csum_cong1: "p1 =o r1 \<Longrightarrow> p1 +c q =o r1 +c q"  | 
|
| 76951 | 183  | 
by (simp only: csum_def ordIso_Plus_cong1)  | 
| 54474 | 184  | 
|
185  | 
lemma csum_cong2: "p2 =o r2 \<Longrightarrow> q +c p2 =o q +c r2"  | 
|
| 76951 | 186  | 
by (simp only: csum_def ordIso_Plus_cong2)  | 
| 54474 | 187  | 
|
188  | 
lemma csum_mono: "\<lbrakk>p1 \<le>o r1; p2 \<le>o r2\<rbrakk> \<Longrightarrow> p1 +c p2 \<le>o r1 +c r2"  | 
|
| 76951 | 189  | 
by (simp only: csum_def ordLeq_Plus_mono)  | 
| 54474 | 190  | 
|
191  | 
lemma csum_mono1: "p1 \<le>o r1 \<Longrightarrow> p1 +c q \<le>o r1 +c q"  | 
|
| 76951 | 192  | 
by (simp only: csum_def ordLeq_Plus_mono1)  | 
| 54474 | 193  | 
|
194  | 
lemma csum_mono2: "p2 \<le>o r2 \<Longrightarrow> q +c p2 \<le>o q +c r2"  | 
|
| 76951 | 195  | 
by (simp only: csum_def ordLeq_Plus_mono2)  | 
| 54474 | 196  | 
|
197  | 
lemma ordLeq_csum1: "Card_order p1 \<Longrightarrow> p1 \<le>o p1 +c p2"  | 
|
| 76951 | 198  | 
by (simp only: csum_def Card_order_Plus1)  | 
| 54474 | 199  | 
|
200  | 
lemma ordLeq_csum2: "Card_order p2 \<Longrightarrow> p2 \<le>o p1 +c p2"  | 
|
| 76951 | 201  | 
by (simp only: csum_def Card_order_Plus2)  | 
| 54474 | 202  | 
|
203  | 
lemma csum_com: "p1 +c p2 =o p2 +c p1"  | 
|
| 76951 | 204  | 
by (simp only: csum_def card_of_Plus_commute)  | 
| 54474 | 205  | 
|
206  | 
lemma csum_assoc: "(p1 +c p2) +c p3 =o p1 +c p2 +c p3"  | 
|
| 76951 | 207  | 
by (simp only: csum_def Field_card_of card_of_Plus_assoc)  | 
| 54474 | 208  | 
|
209  | 
lemma Cfinite_csum: "\<lbrakk>Cfinite r; Cfinite s\<rbrakk> \<Longrightarrow> Cfinite (r +c s)"  | 
|
210  | 
unfolding cfinite_def csum_def Field_card_of using card_of_card_order_on by simp  | 
|
211  | 
||
212  | 
lemma csum_csum: "(r1 +c r2) +c (r3 +c r4) =o (r1 +c r3) +c (r2 +c r4)"  | 
|
213  | 
proof -  | 
|
214  | 
have "(r1 +c r2) +c (r3 +c r4) =o r1 +c r2 +c (r3 +c r4)"  | 
|
| 55811 | 215  | 
by (rule csum_assoc)  | 
| 54474 | 216  | 
also have "r1 +c r2 +c (r3 +c r4) =o r1 +c (r2 +c r3) +c r4"  | 
| 55811 | 217  | 
by (intro csum_assoc csum_cong2 ordIso_symmetric)  | 
| 54474 | 218  | 
also have "r1 +c (r2 +c r3) +c r4 =o r1 +c (r3 +c r2) +c r4"  | 
| 55811 | 219  | 
by (intro csum_com csum_cong1 csum_cong2)  | 
| 54474 | 220  | 
also have "r1 +c (r3 +c r2) +c r4 =o r1 +c r3 +c r2 +c r4"  | 
| 55811 | 221  | 
by (intro csum_assoc csum_cong2 ordIso_symmetric)  | 
| 54474 | 222  | 
also have "r1 +c r3 +c r2 +c r4 =o (r1 +c r3) +c (r2 +c r4)"  | 
| 55811 | 223  | 
by (intro csum_assoc ordIso_symmetric)  | 
| 54474 | 224  | 
finally show ?thesis .  | 
225  | 
qed  | 
|
226  | 
||
227  | 
lemma Plus_csum: "|A <+> B| =o |A| +c |B|"  | 
|
| 76951 | 228  | 
by (simp only: csum_def Field_card_of card_of_refl)  | 
| 54474 | 229  | 
|
230  | 
lemma Un_csum: "|A \<union> B| \<le>o |A| +c |B|"  | 
|
| 76951 | 231  | 
using ordLeq_ordIso_trans[OF card_of_Un_Plus_ordLeq Plus_csum] by blast  | 
| 54474 | 232  | 
|
| 60758 | 233  | 
subsection \<open>One\<close>  | 
| 54474 | 234  | 
|
235  | 
definition cone where  | 
|
236  | 
  "cone = card_of {()}"
 | 
|
237  | 
||
238  | 
lemma Card_order_cone: "Card_order cone"  | 
|
| 76951 | 239  | 
unfolding cone_def by (rule card_of_Card_order)  | 
| 54474 | 240  | 
|
241  | 
lemma Cfinite_cone: "Cfinite cone"  | 
|
242  | 
unfolding cfinite_def by (simp add: Card_order_cone)  | 
|
243  | 
||
244  | 
lemma cone_not_czero: "\<not> (cone =o czero)"  | 
|
| 76951 | 245  | 
unfolding czero_def cone_def ordIso_iff_ordLeq  | 
246  | 
using card_of_empty3 empty_not_insert by blast  | 
|
| 54474 | 247  | 
|
248  | 
lemma cone_ordLeq_Cnotzero: "Cnotzero r \<Longrightarrow> cone \<le>o r"  | 
|
| 76951 | 249  | 
unfolding cone_def by (rule Card_order_singl_ordLeq) (auto intro: czeroI)  | 
| 54474 | 250  | 
|
251  | 
||
| 60758 | 252  | 
subsection \<open>Two\<close>  | 
| 54474 | 253  | 
|
254  | 
definition ctwo where  | 
|
255  | 
"ctwo = |UNIV :: bool set|"  | 
|
256  | 
||
257  | 
lemma Card_order_ctwo: "Card_order ctwo"  | 
|
| 76951 | 258  | 
unfolding ctwo_def by (rule card_of_Card_order)  | 
| 54474 | 259  | 
|
260  | 
lemma ctwo_not_czero: "\<not> (ctwo =o czero)"  | 
|
| 76951 | 261  | 
using card_of_empty3[of "UNIV :: bool set"] ordIso_iff_ordLeq  | 
262  | 
unfolding czero_def ctwo_def using UNIV_not_empty by auto  | 
|
| 54474 | 263  | 
|
264  | 
lemma ctwo_Cnotzero: "Cnotzero ctwo"  | 
|
| 76951 | 265  | 
by (simp add: ctwo_not_czero Card_order_ctwo)  | 
| 54474 | 266  | 
|
267  | 
||
| 60758 | 268  | 
subsection \<open>Family sum\<close>  | 
| 54474 | 269  | 
|
270  | 
definition Csum where  | 
|
271  | 
"Csum r rs \<equiv> |SIGMA i : Field r. Field (rs i)|"  | 
|
272  | 
||
273  | 
(* Similar setup to the one for SIGMA from theory Big_Operators: *)  | 
|
274  | 
syntax "_Csum" ::  | 
|
275  | 
  "pttrn => ('a * 'a) set => 'b * 'b set => (('a * 'b) * ('a * 'b)) set"
 | 
|
276  | 
  ("(3CSUM _:_. _)" [0, 51, 10] 10)
 | 
|
277  | 
||
278  | 
translations  | 
|
279  | 
"CSUM i:r. rs" == "CONST Csum r (%i. rs)"  | 
|
280  | 
||
281  | 
lemma SIGMA_CSUM: "|SIGMA i : I. As i| = (CSUM i : |I|. |As i| )"  | 
|
| 76951 | 282  | 
by (auto simp: Csum_def Field_card_of)  | 
| 54474 | 283  | 
|
284  | 
(* NB: Always, under the cardinal operator,  | 
|
285  | 
operations on sets are reduced automatically to operations on cardinals.  | 
|
286  | 
This should make cardinal reasoning more direct and natural. *)  | 
|
287  | 
||
288  | 
||
| 60758 | 289  | 
subsection \<open>Product\<close>  | 
| 54474 | 290  | 
|
291  | 
definition cprod (infixr "*c" 80) where  | 
|
| 61943 | 292  | 
"r1 *c r2 = |Field r1 \<times> Field r2|"  | 
| 54474 | 293  | 
|
294  | 
lemma card_order_cprod:  | 
|
295  | 
assumes "card_order r1" "card_order r2"  | 
|
296  | 
shows "card_order (r1 *c r2)"  | 
|
297  | 
proof -  | 
|
| 76951 | 298  | 
have "Field r1 = UNIV" "Field r2 = UNIV"  | 
299  | 
using assms card_order_on_Card_order by auto  | 
|
| 54474 | 300  | 
thus ?thesis by (auto simp: cprod_def card_of_card_order_on)  | 
301  | 
qed  | 
|
302  | 
||
303  | 
lemma Card_order_cprod: "Card_order (r1 *c r2)"  | 
|
| 76951 | 304  | 
by (simp only: cprod_def Field_card_of card_of_card_order_on)  | 
| 54474 | 305  | 
|
306  | 
lemma cprod_mono1: "p1 \<le>o r1 \<Longrightarrow> p1 *c q \<le>o r1 *c q"  | 
|
| 76951 | 307  | 
by (simp only: cprod_def ordLeq_Times_mono1)  | 
| 54474 | 308  | 
|
309  | 
lemma cprod_mono2: "p2 \<le>o r2 \<Longrightarrow> q *c p2 \<le>o q *c r2"  | 
|
| 76951 | 310  | 
by (simp only: cprod_def ordLeq_Times_mono2)  | 
| 54474 | 311  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
312  | 
lemma cprod_mono: "\<lbrakk>p1 \<le>o r1; p2 \<le>o r2\<rbrakk> \<Longrightarrow> p1 *c p2 \<le>o r1 *c r2"  | 
| 76951 | 313  | 
by (rule ordLeq_transitive[OF cprod_mono1 cprod_mono2])  | 
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
314  | 
|
| 54474 | 315  | 
lemma ordLeq_cprod2: "\<lbrakk>Cnotzero p1; Card_order p2\<rbrakk> \<Longrightarrow> p2 \<le>o p1 *c p2"  | 
| 76951 | 316  | 
unfolding cprod_def by (rule Card_order_Times2) (auto intro: czeroI)  | 
| 54474 | 317  | 
|
318  | 
lemma cinfinite_cprod: "\<lbrakk>cinfinite r1; cinfinite r2\<rbrakk> \<Longrightarrow> cinfinite (r1 *c r2)"  | 
|
| 76951 | 319  | 
by (simp add: cinfinite_def cprod_def Field_card_of infinite_cartesian_product)  | 
| 54474 | 320  | 
|
321  | 
lemma cinfinite_cprod2: "\<lbrakk>Cnotzero r1; Cinfinite r2\<rbrakk> \<Longrightarrow> cinfinite (r1 *c r2)"  | 
|
| 76951 | 322  | 
by (rule cinfinite_mono) (auto intro: ordLeq_cprod2)  | 
| 54474 | 323  | 
|
324  | 
lemma Cinfinite_cprod2: "\<lbrakk>Cnotzero r1; Cinfinite r2\<rbrakk> \<Longrightarrow> Cinfinite (r1 *c r2)"  | 
|
| 76951 | 325  | 
by (blast intro: cinfinite_cprod2 Card_order_cprod)  | 
| 54474 | 326  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
327  | 
lemma cprod_cong: "\<lbrakk>p1 =o r1; p2 =o r2\<rbrakk> \<Longrightarrow> p1 *c p2 =o r1 *c r2"  | 
| 76951 | 328  | 
unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono)  | 
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
329  | 
|
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
330  | 
lemma cprod_cong1: "\<lbrakk>p1 =o r1\<rbrakk> \<Longrightarrow> p1 *c p2 =o r1 *c p2"  | 
| 76951 | 331  | 
unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono1)  | 
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
332  | 
|
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
333  | 
lemma cprod_cong2: "p2 =o r2 \<Longrightarrow> q *c p2 =o q *c r2"  | 
| 76951 | 334  | 
unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono2)  | 
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
335  | 
|
| 54474 | 336  | 
lemma cprod_com: "p1 *c p2 =o p2 *c p1"  | 
| 76951 | 337  | 
by (simp only: cprod_def card_of_Times_commute)  | 
| 54474 | 338  | 
|
339  | 
lemma card_of_Csum_Times:  | 
|
340  | 
"\<forall>i \<in> I. |A i| \<le>o |B| \<Longrightarrow> (CSUM i : |I|. |A i| ) \<le>o |I| *c |B|"  | 
|
| 76951 | 341  | 
by (simp only: Csum_def cprod_def Field_card_of card_of_Sigma_mono1)  | 
| 54474 | 342  | 
|
343  | 
lemma card_of_Csum_Times':  | 
|
344  | 
assumes "Card_order r" "\<forall>i \<in> I. |A i| \<le>o r"  | 
|
345  | 
shows "(CSUM i : |I|. |A i| ) \<le>o |I| *c r"  | 
|
346  | 
proof -  | 
|
347  | 
from assms(1) have *: "r =o |Field r|" by (simp add: card_of_unique)  | 
|
348  | 
with assms(2) have "\<forall>i \<in> I. |A i| \<le>o |Field r|" by (blast intro: ordLeq_ordIso_trans)  | 
|
349  | 
hence "(CSUM i : |I|. |A i| ) \<le>o |I| *c |Field r|" by (simp only: card_of_Csum_Times)  | 
|
350  | 
also from * have "|I| *c |Field r| \<le>o |I| *c r"  | 
|
351  | 
by (simp only: Field_card_of card_of_refl cprod_def ordIso_imp_ordLeq)  | 
|
352  | 
finally show ?thesis .  | 
|
353  | 
qed  | 
|
354  | 
||
355  | 
lemma cprod_csum_distrib1: "r1 *c r2 +c r1 *c r3 =o r1 *c (r2 +c r3)"  | 
|
| 76951 | 356  | 
unfolding csum_def cprod_def by (simp add: Field_card_of card_of_Times_Plus_distrib ordIso_symmetric)  | 
| 54474 | 357  | 
|
358  | 
lemma csum_absorb2': "\<lbrakk>Card_order r2; r1 \<le>o r2; cinfinite r1 \<or> cinfinite r2\<rbrakk> \<Longrightarrow> r1 +c r2 =o r2"  | 
|
| 76951 | 359  | 
unfolding csum_def  | 
360  | 
using Card_order_Plus_infinite  | 
|
361  | 
by (fastforce simp: cinfinite_def dest: cinfinite_mono)  | 
|
| 54474 | 362  | 
|
363  | 
lemma csum_absorb1':  | 
|
364  | 
assumes card: "Card_order r2"  | 
|
| 76951 | 365  | 
and r12: "r1 \<le>o r2" and cr12: "cinfinite r1 \<or> cinfinite r2"  | 
| 54474 | 366  | 
shows "r2 +c r1 =o r2"  | 
| 76951 | 367  | 
proof -  | 
368  | 
have "r1 +c r2 =o r2"  | 
|
369  | 
by (simp add: csum_absorb2' assms)  | 
|
370  | 
then show ?thesis  | 
|
371  | 
by (blast intro: ordIso_transitive csum_com)  | 
|
372  | 
qed  | 
|
| 54474 | 373  | 
|
374  | 
lemma csum_absorb1: "\<lbrakk>Cinfinite r2; r1 \<le>o r2\<rbrakk> \<Longrightarrow> r2 +c r1 =o r2"  | 
|
| 76951 | 375  | 
by (rule csum_absorb1') auto  | 
| 54474 | 376  | 
|
| 75624 | 377  | 
lemma csum_absorb2: "\<lbrakk>Cinfinite r2 ; r1 \<le>o r2\<rbrakk> \<Longrightarrow> r1 +c r2 =o r2"  | 
378  | 
using ordIso_transitive csum_com csum_absorb1 by blast  | 
|
379  | 
||
380  | 
lemma regularCard_csum:  | 
|
381  | 
assumes "Cinfinite r" "Cinfinite s" "regularCard r" "regularCard s"  | 
|
| 76951 | 382  | 
shows "regularCard (r +c s)"  | 
| 75624 | 383  | 
proof (cases "r \<le>o s")  | 
384  | 
case True  | 
|
385  | 
then show ?thesis using regularCard_ordIso[of s] csum_absorb2'[THEN ordIso_symmetric] assms by auto  | 
|
386  | 
next  | 
|
387  | 
case False  | 
|
388  | 
have "Well_order s" "Well_order r" using assms card_order_on_well_order_on by auto  | 
|
389  | 
then have "s \<le>o r" using not_ordLeq_iff_ordLess False ordLess_imp_ordLeq by auto  | 
|
390  | 
then show ?thesis using regularCard_ordIso[of r] csum_absorb1'[THEN ordIso_symmetric] assms by auto  | 
|
391  | 
qed  | 
|
392  | 
||
393  | 
lemma csum_mono_strict:  | 
|
394  | 
assumes Card_order: "Card_order r" "Card_order q"  | 
|
| 76951 | 395  | 
and Cinfinite: "Cinfinite r'" "Cinfinite q'"  | 
396  | 
and less: "r <o r'" "q <o q'"  | 
|
397  | 
shows "r +c q <o r' +c q'"  | 
|
| 75624 | 398  | 
proof -  | 
399  | 
have Well_order: "Well_order r" "Well_order q" "Well_order r'" "Well_order q'"  | 
|
400  | 
using card_order_on_well_order_on Card_order Cinfinite by auto  | 
|
401  | 
show ?thesis  | 
|
402  | 
proof (cases "Cinfinite r")  | 
|
403  | 
case outer: True  | 
|
404  | 
then show ?thesis  | 
|
405  | 
proof (cases "Cinfinite q")  | 
|
406  | 
case inner: True  | 
|
407  | 
then show ?thesis  | 
|
408  | 
proof (cases "r \<le>o q")  | 
|
409  | 
case True  | 
|
410  | 
then have "r +c q =o q" using csum_absorb2 inner by blast  | 
|
411  | 
then show ?thesis  | 
|
412  | 
using ordIso_ordLess_trans ordLess_ordLeq_trans less Cinfinite ordLeq_csum2 by blast  | 
|
413  | 
next  | 
|
414  | 
case False  | 
|
415  | 
then have "q \<le>o r" using not_ordLeq_iff_ordLess Well_order ordLess_imp_ordLeq by blast  | 
|
416  | 
then have "r +c q =o r" using csum_absorb1 outer by blast  | 
|
417  | 
then show ?thesis  | 
|
418  | 
using ordIso_ordLess_trans ordLess_ordLeq_trans less Cinfinite ordLeq_csum1 by blast  | 
|
419  | 
qed  | 
|
420  | 
next  | 
|
421  | 
case False  | 
|
422  | 
then have "Cfinite q" using Card_order cinfinite_def cfinite_def by blast  | 
|
423  | 
then have "q \<le>o r" using finite_ordLess_infinite cfinite_def cinfinite_def outer  | 
|
424  | 
Well_order ordLess_imp_ordLeq by blast  | 
|
425  | 
then have "r +c q =o r" by (rule csum_absorb1[OF outer])  | 
|
426  | 
then show ?thesis using ordIso_ordLess_trans ordLess_ordLeq_trans less ordLeq_csum1 Cinfinite by blast  | 
|
427  | 
qed  | 
|
428  | 
next  | 
|
429  | 
case False  | 
|
430  | 
then have outer: "Cfinite r" using Card_order cinfinite_def cfinite_def by blast  | 
|
431  | 
then show ?thesis  | 
|
432  | 
proof (cases "Cinfinite q")  | 
|
433  | 
case True  | 
|
434  | 
then have "r \<le>o q" using finite_ordLess_infinite cinfinite_def cfinite_def outer Well_order  | 
|
| 76951 | 435  | 
ordLess_imp_ordLeq by blast  | 
| 75624 | 436  | 
then have "r +c q =o q" by (rule csum_absorb2[OF True])  | 
437  | 
then show ?thesis using ordIso_ordLess_trans ordLess_ordLeq_trans less ordLeq_csum2 Cinfinite by blast  | 
|
438  | 
next  | 
|
439  | 
case False  | 
|
440  | 
then have "Cfinite q" using Card_order cinfinite_def cfinite_def by blast  | 
|
441  | 
then have "Cfinite (r +c q)" using Cfinite_csum outer by blast  | 
|
442  | 
moreover have "Cinfinite (r' +c q')" using Cinfinite_csum1 Cinfinite by blast  | 
|
443  | 
ultimately show ?thesis using Cfinite_ordLess_Cinfinite by blast  | 
|
444  | 
qed  | 
|
445  | 
qed  | 
|
446  | 
qed  | 
|
| 54474 | 447  | 
|
| 60758 | 448  | 
subsection \<open>Exponentiation\<close>  | 
| 54474 | 449  | 
|
450  | 
definition cexp (infixr "^c" 90) where  | 
|
451  | 
"r1 ^c r2 \<equiv> |Func (Field r2) (Field r1)|"  | 
|
452  | 
||
453  | 
lemma Card_order_cexp: "Card_order (r1 ^c r2)"  | 
|
| 76951 | 454  | 
unfolding cexp_def by (rule card_of_Card_order)  | 
| 54474 | 455  | 
|
456  | 
lemma cexp_mono':  | 
|
457  | 
assumes 1: "p1 \<le>o r1" and 2: "p2 \<le>o r2"  | 
|
| 76951 | 458  | 
    and n: "Field p2 = {} \<Longrightarrow> Field r2 = {}"
 | 
| 54474 | 459  | 
shows "p1 ^c p2 \<le>o r1 ^c r2"  | 
460  | 
proof(cases "Field p1 = {}")
 | 
|
461  | 
case True  | 
|
| 55811 | 462  | 
  hence "Field p2 \<noteq> {} \<Longrightarrow> Func (Field p2) {} = {}" unfolding Func_is_emp by simp
 | 
463  | 
with True have "|Field |Func (Field p2) (Field p1)|| \<le>o cone"  | 
|
| 54474 | 464  | 
unfolding cone_def Field_card_of  | 
| 55811 | 465  | 
    by (cases "Field p2 = {}", auto intro: surj_imp_ordLeq simp: Func_empty)
 | 
| 54474 | 466  | 
hence "|Func (Field p2) (Field p1)| \<le>o cone" by (simp add: Field_card_of cexp_def)  | 
467  | 
hence "p1 ^c p2 \<le>o cone" unfolding cexp_def .  | 
|
468  | 
thus ?thesis  | 
|
469  | 
  proof (cases "Field p2 = {}")
 | 
|
470  | 
case True  | 
|
471  | 
    with n have "Field r2 = {}" .
 | 
|
| 55604 | 472  | 
hence "cone \<le>o r1 ^c r2" unfolding cone_def cexp_def Func_def  | 
473  | 
by (auto intro: card_of_ordLeqI[where f="\<lambda>_ _. undefined"])  | 
|
| 60758 | 474  | 
thus ?thesis using \<open>p1 ^c p2 \<le>o cone\<close> ordLeq_transitive by auto  | 
| 54474 | 475  | 
next  | 
476  | 
case False with True have "|Field (p1 ^c p2)| =o czero"  | 
|
477  | 
unfolding card_of_ordIso_czero_iff_empty cexp_def Field_card_of Func_def by auto  | 
|
478  | 
thus ?thesis unfolding cexp_def card_of_ordIso_czero_iff_empty Field_card_of  | 
|
479  | 
by (simp add: card_of_empty)  | 
|
480  | 
qed  | 
|
481  | 
next  | 
|
482  | 
case False  | 
|
483  | 
have 1: "|Field p1| \<le>o |Field r1|" and 2: "|Field p2| \<le>o |Field r2|"  | 
|
484  | 
using 1 2 by (auto simp: card_of_mono2)  | 
|
485  | 
obtain f1 where f1: "f1 ` Field r1 = Field p1"  | 
|
486  | 
using 1 unfolding card_of_ordLeq2[OF False, symmetric] by auto  | 
|
487  | 
obtain f2 where f2: "inj_on f2 (Field p2)" "f2 ` Field p2 \<subseteq> Field r2"  | 
|
488  | 
using 2 unfolding card_of_ordLeq[symmetric] by blast  | 
|
489  | 
have 0: "Func_map (Field p2) f1 f2 ` (Field (r1 ^c r2)) = Field (p1 ^c p2)"  | 
|
490  | 
unfolding cexp_def Field_card_of using Func_map_surj[OF f1 f2 n, symmetric] .  | 
|
491  | 
  have 00: "Field (p1 ^c p2) \<noteq> {}" unfolding cexp_def Field_card_of Func_is_emp
 | 
|
492  | 
using False by simp  | 
|
493  | 
show ?thesis  | 
|
494  | 
using 0 card_of_ordLeq2[OF 00] unfolding cexp_def Field_card_of by blast  | 
|
495  | 
qed  | 
|
496  | 
||
497  | 
lemma cexp_mono:  | 
|
498  | 
assumes 1: "p1 \<le>o r1" and 2: "p2 \<le>o r2"  | 
|
| 76951 | 499  | 
and n: "p2 =o czero \<Longrightarrow> r2 =o czero" and card: "Card_order p2"  | 
| 54474 | 500  | 
shows "p1 ^c p2 \<le>o r1 ^c r2"  | 
| 55811 | 501  | 
by (rule cexp_mono'[OF 1 2 czeroE[OF n[OF czeroI[OF card]]]])  | 
| 54474 | 502  | 
|
503  | 
lemma cexp_mono1:  | 
|
504  | 
assumes 1: "p1 \<le>o r1" and q: "Card_order q"  | 
|
505  | 
shows "p1 ^c q \<le>o r1 ^c q"  | 
|
| 76951 | 506  | 
using ordLeq_refl[OF q] by (rule cexp_mono[OF 1]) (auto simp: q)  | 
| 54474 | 507  | 
|
508  | 
lemma cexp_mono2':  | 
|
509  | 
assumes 2: "p2 \<le>o r2" and q: "Card_order q"  | 
|
| 76951 | 510  | 
    and n: "Field p2 = {} \<Longrightarrow> Field r2 = {}"
 | 
| 54474 | 511  | 
shows "q ^c p2 \<le>o q ^c r2"  | 
| 76951 | 512  | 
using ordLeq_refl[OF q] by (rule cexp_mono'[OF _ 2 n]) auto  | 
| 54474 | 513  | 
|
514  | 
lemma cexp_mono2:  | 
|
515  | 
assumes 2: "p2 \<le>o r2" and q: "Card_order q"  | 
|
| 76951 | 516  | 
and n: "p2 =o czero \<Longrightarrow> r2 =o czero" and card: "Card_order p2"  | 
| 54474 | 517  | 
shows "q ^c p2 \<le>o q ^c r2"  | 
| 76951 | 518  | 
using ordLeq_refl[OF q] by (rule cexp_mono[OF _ 2 n card]) auto  | 
| 54474 | 519  | 
|
520  | 
lemma cexp_mono2_Cnotzero:  | 
|
521  | 
assumes "p2 \<le>o r2" "Card_order q" "Cnotzero p2"  | 
|
522  | 
shows "q ^c p2 \<le>o q ^c r2"  | 
|
| 76951 | 523  | 
using assms(3) czeroI by (blast intro: cexp_mono2'[OF assms(1,2)])  | 
| 54474 | 524  | 
|
525  | 
lemma cexp_cong:  | 
|
526  | 
assumes 1: "p1 =o r1" and 2: "p2 =o r2"  | 
|
| 76951 | 527  | 
and Cr: "Card_order r2"  | 
528  | 
and Cp: "Card_order p2"  | 
|
| 54474 | 529  | 
shows "p1 ^c p2 =o r1 ^c r2"  | 
530  | 
proof -  | 
|
531  | 
obtain f where "bij_betw f (Field p2) (Field r2)"  | 
|
532  | 
using 2 card_of_ordIso[of "Field p2" "Field r2"] card_of_cong by auto  | 
|
533  | 
  hence 0: "Field p2 = {} \<longleftrightarrow> Field r2 = {}" unfolding bij_betw_def by auto
 | 
|
534  | 
have r: "p2 =o czero \<Longrightarrow> r2 =o czero"  | 
|
535  | 
and p: "r2 =o czero \<Longrightarrow> p2 =o czero"  | 
|
| 76951 | 536  | 
using 0 Cr Cp czeroE czeroI by auto  | 
| 54474 | 537  | 
show ?thesis using 0 1 2 unfolding ordIso_iff_ordLeq  | 
| 55811 | 538  | 
using r p cexp_mono[OF _ _ _ Cp] cexp_mono[OF _ _ _ Cr] by blast  | 
| 54474 | 539  | 
qed  | 
540  | 
||
541  | 
lemma cexp_cong1:  | 
|
542  | 
assumes 1: "p1 =o r1" and q: "Card_order q"  | 
|
543  | 
shows "p1 ^c q =o r1 ^c q"  | 
|
| 76951 | 544  | 
by (rule cexp_cong[OF 1 _ q q]) (rule ordIso_refl[OF q])  | 
| 54474 | 545  | 
|
546  | 
lemma cexp_cong2:  | 
|
547  | 
assumes 2: "p2 =o r2" and q: "Card_order q" and p: "Card_order p2"  | 
|
548  | 
shows "q ^c p2 =o q ^c r2"  | 
|
| 76951 | 549  | 
by (rule cexp_cong[OF _ 2]) (auto simp only: ordIso_refl Card_order_ordIso2[OF p 2] q p)  | 
| 54474 | 550  | 
|
551  | 
lemma cexp_cone:  | 
|
552  | 
assumes "Card_order r"  | 
|
553  | 
shows "r ^c cone =o r"  | 
|
554  | 
proof -  | 
|
555  | 
have "r ^c cone =o |Field r|"  | 
|
556  | 
unfolding cexp_def cone_def Field_card_of Func_empty  | 
|
557  | 
card_of_ordIso[symmetric] bij_betw_def Func_def inj_on_def image_def  | 
|
558  | 
by (rule exI[of _ "\<lambda>f. f ()"]) auto  | 
|
559  | 
also have "|Field r| =o r" by (rule card_of_Field_ordIso[OF assms])  | 
|
560  | 
finally show ?thesis .  | 
|
561  | 
qed  | 
|
562  | 
||
563  | 
lemma cexp_cprod:  | 
|
564  | 
assumes r1: "Card_order r1"  | 
|
565  | 
shows "(r1 ^c r2) ^c r3 =o r1 ^c (r2 *c r3)" (is "?L =o ?R")  | 
|
566  | 
proof -  | 
|
567  | 
have "?L =o r1 ^c (r3 *c r2)"  | 
|
568  | 
unfolding cprod_def cexp_def Field_card_of  | 
|
569  | 
using card_of_Func_Times by(rule ordIso_symmetric)  | 
|
570  | 
also have "r1 ^c (r3 *c r2) =o ?R"  | 
|
| 76951 | 571  | 
using cprod_com r1 by (intro cexp_cong2, auto simp: Card_order_cprod)  | 
| 54474 | 572  | 
finally show ?thesis .  | 
573  | 
qed  | 
|
574  | 
||
575  | 
lemma cprod_infinite1': "\<lbrakk>Cinfinite r; Cnotzero p; p \<le>o r\<rbrakk> \<Longrightarrow> r *c p =o r"  | 
|
| 76951 | 576  | 
unfolding cinfinite_def cprod_def  | 
577  | 
by (rule Card_order_Times_infinite[THEN conjunct1]) (blast intro: czeroI)+  | 
|
| 54474 | 578  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
579  | 
lemma cprod_infinite: "Cinfinite r \<Longrightarrow> r *c r =o r"  | 
| 76951 | 580  | 
using cprod_infinite1' Cinfinite_Cnotzero ordLeq_refl by blast  | 
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
581  | 
|
| 54474 | 582  | 
lemma cexp_cprod_ordLeq:  | 
583  | 
assumes r1: "Card_order r1" and r2: "Cinfinite r2"  | 
|
| 76951 | 584  | 
and r3: "Cnotzero r3" "r3 \<le>o r2"  | 
| 54474 | 585  | 
shows "(r1 ^c r2) ^c r3 =o r1 ^c r2" (is "?L =o ?R")  | 
586  | 
proof-  | 
|
587  | 
have "?L =o r1 ^c (r2 *c r3)" using cexp_cprod[OF r1] .  | 
|
588  | 
also have "r1 ^c (r2 *c r3) =o ?R"  | 
|
| 76951 | 589  | 
using assms by (fastforce simp: Card_order_cprod intro: cprod_infinite1' cexp_cong2)  | 
| 54474 | 590  | 
finally show ?thesis .  | 
591  | 
qed  | 
|
592  | 
||
593  | 
lemma Cnotzero_UNIV: "Cnotzero |UNIV|"  | 
|
| 76951 | 594  | 
by (auto simp: card_of_Card_order card_of_ordIso_czero_iff_empty)  | 
| 54474 | 595  | 
|
596  | 
lemma ordLess_ctwo_cexp:  | 
|
597  | 
assumes "Card_order r"  | 
|
598  | 
shows "r <o ctwo ^c r"  | 
|
599  | 
proof -  | 
|
600  | 
have "r <o |Pow (Field r)|" using assms by (rule Card_order_Pow)  | 
|
601  | 
also have "|Pow (Field r)| =o ctwo ^c r"  | 
|
602  | 
unfolding ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)  | 
|
603  | 
finally show ?thesis .  | 
|
604  | 
qed  | 
|
605  | 
||
606  | 
lemma ordLeq_cexp1:  | 
|
607  | 
assumes "Cnotzero r" "Card_order q"  | 
|
608  | 
shows "q \<le>o q ^c r"  | 
|
609  | 
proof (cases "q =o (czero :: 'a rel)")  | 
|
610  | 
case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans)  | 
|
611  | 
next  | 
|
612  | 
case False  | 
|
| 76951 | 613  | 
have "q =o q ^c cone"  | 
614  | 
by (blast intro: assms ordIso_symmetric cexp_cone)  | 
|
615  | 
also have "q ^c cone \<le>o q ^c r"  | 
|
616  | 
using assms  | 
|
617  | 
by (intro cexp_mono2) (simp_all add: cone_ordLeq_Cnotzero cone_not_czero Card_order_cone)  | 
|
618  | 
finally show ?thesis .  | 
|
| 54474 | 619  | 
qed  | 
620  | 
||
621  | 
lemma ordLeq_cexp2:  | 
|
622  | 
assumes "ctwo \<le>o q" "Card_order r"  | 
|
623  | 
shows "r \<le>o q ^c r"  | 
|
624  | 
proof (cases "r =o (czero :: 'a rel)")  | 
|
625  | 
case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans)  | 
|
626  | 
next  | 
|
| 76951 | 627  | 
case False  | 
628  | 
have "r <o ctwo ^c r"  | 
|
629  | 
by (blast intro: assms ordLess_ctwo_cexp)  | 
|
630  | 
also have "ctwo ^c r \<le>o q ^c r"  | 
|
631  | 
by (blast intro: assms cexp_mono1)  | 
|
632  | 
finally show ?thesis by (rule ordLess_imp_ordLeq)  | 
|
| 54474 | 633  | 
qed  | 
634  | 
||
635  | 
lemma cinfinite_cexp: "\<lbrakk>ctwo \<le>o q; Cinfinite r\<rbrakk> \<Longrightarrow> cinfinite (q ^c r)"  | 
|
| 76951 | 636  | 
by (rule cinfinite_mono[OF ordLeq_cexp2]) simp_all  | 
| 54474 | 637  | 
|
638  | 
lemma Cinfinite_cexp:  | 
|
639  | 
"\<lbrakk>ctwo \<le>o q; Cinfinite r\<rbrakk> \<Longrightarrow> Cinfinite (q ^c r)"  | 
|
| 76951 | 640  | 
by (simp add: cinfinite_cexp Card_order_cexp)  | 
| 54474 | 641  | 
|
| 75624 | 642  | 
lemma card_order_cexp:  | 
643  | 
assumes "card_order r1" "card_order r2"  | 
|
644  | 
shows "card_order (r1 ^c r2)"  | 
|
645  | 
proof -  | 
|
646  | 
have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto  | 
|
647  | 
thus ?thesis unfolding cexp_def Func_def using card_of_card_order_on by simp  | 
|
648  | 
qed  | 
|
649  | 
||
| 54474 | 650  | 
lemma ctwo_ordLess_natLeq: "ctwo <o natLeq"  | 
| 76951 | 651  | 
unfolding ctwo_def using finite_UNIV natLeq_cinfinite natLeq_Card_order  | 
652  | 
by (intro Cfinite_ordLess_Cinfinite) (auto simp: cfinite_def card_of_Card_order)  | 
|
| 54474 | 653  | 
|
654  | 
lemma ctwo_ordLess_Cinfinite: "Cinfinite r \<Longrightarrow> ctwo <o r"  | 
|
| 76951 | 655  | 
by (rule ordLess_ordLeq_trans[OF ctwo_ordLess_natLeq natLeq_ordLeq_cinfinite])  | 
| 54474 | 656  | 
|
657  | 
lemma ctwo_ordLeq_Cinfinite:  | 
|
658  | 
assumes "Cinfinite r"  | 
|
659  | 
shows "ctwo \<le>o r"  | 
|
| 76951 | 660  | 
by (rule ordLess_imp_ordLeq[OF ctwo_ordLess_Cinfinite[OF assms]])  | 
| 54474 | 661  | 
|
662  | 
lemma Un_Cinfinite_bound: "\<lbrakk>|A| \<le>o r; |B| \<le>o r; Cinfinite r\<rbrakk> \<Longrightarrow> |A \<union> B| \<le>o r"  | 
|
| 76951 | 663  | 
by (auto simp add: cinfinite_def card_of_Un_ordLeq_infinite_Field)  | 
| 54474 | 664  | 
|
| 75624 | 665  | 
lemma Un_Cinfinite_bound_strict: "\<lbrakk>|A| <o r; |B| <o r; Cinfinite r\<rbrakk> \<Longrightarrow> |A \<union> B| <o r"  | 
| 76951 | 666  | 
by (auto simp add: cinfinite_def card_of_Un_ordLess_infinite_Field)  | 
| 75624 | 667  | 
|
| 54474 | 668  | 
lemma UNION_Cinfinite_bound: "\<lbrakk>|I| \<le>o r; \<forall>i \<in> I. |A i| \<le>o r; Cinfinite r\<rbrakk> \<Longrightarrow> |\<Union>i \<in> I. A i| \<le>o r"  | 
| 76951 | 669  | 
by (auto simp add: card_of_UNION_ordLeq_infinite_Field cinfinite_def)  | 
| 54474 | 670  | 
|
671  | 
lemma csum_cinfinite_bound:  | 
|
672  | 
assumes "p \<le>o r" "q \<le>o r" "Card_order p" "Card_order q" "Cinfinite r"  | 
|
673  | 
shows "p +c q \<le>o r"  | 
|
674  | 
proof -  | 
|
| 76951 | 675  | 
have "|Field p| \<le>o r" "|Field q| \<le>o r"  | 
676  | 
using assms card_of_least ordLeq_transitive unfolding card_order_on_def by blast+  | 
|
| 54474 | 677  | 
with assms show ?thesis unfolding cinfinite_def csum_def  | 
678  | 
by (blast intro: card_of_Plus_ordLeq_infinite_Field)  | 
|
679  | 
qed  | 
|
680  | 
||
681  | 
lemma cprod_cinfinite_bound:  | 
|
682  | 
assumes "p \<le>o r" "q \<le>o r" "Card_order p" "Card_order q" "Cinfinite r"  | 
|
683  | 
shows "p *c q \<le>o r"  | 
|
684  | 
proof -  | 
|
685  | 
from assms(1-4) have "|Field p| \<le>o r" "|Field q| \<le>o r"  | 
|
686  | 
unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+  | 
|
687  | 
with assms show ?thesis unfolding cinfinite_def cprod_def  | 
|
688  | 
by (blast intro: card_of_Times_ordLeq_infinite_Field)  | 
|
689  | 
qed  | 
|
690  | 
||
| 75624 | 691  | 
lemma cprod_infinite2': "\<lbrakk>Cnotzero r1; Cinfinite r2; r1 \<le>o r2\<rbrakk> \<Longrightarrow> r1 *c r2 =o r2"  | 
692  | 
unfolding ordIso_iff_ordLeq  | 
|
693  | 
by (intro conjI cprod_cinfinite_bound ordLeq_cprod2 ordLeq_refl)  | 
|
694  | 
(auto dest!: ordIso_imp_ordLeq not_ordLeq_ordLess simp: czero_def Card_order_empty)  | 
|
695  | 
||
696  | 
lemma regularCard_cprod:  | 
|
697  | 
assumes "Cinfinite r" "Cinfinite s" "regularCard r" "regularCard s"  | 
|
| 76951 | 698  | 
shows "regularCard (r *c s)"  | 
| 75624 | 699  | 
proof (cases "r \<le>o s")  | 
700  | 
case True  | 
|
| 76951 | 701  | 
with assms Cinfinite_Cnotzero show ?thesis  | 
702  | 
by (force intro: regularCard_ordIso ordIso_symmetric[OF cprod_infinite2'])  | 
|
| 75624 | 703  | 
next  | 
704  | 
case False  | 
|
705  | 
have "Well_order r" "Well_order s" using assms card_order_on_well_order_on by auto  | 
|
| 76951 | 706  | 
then have "s \<le>o r" using not_ordLeq_iff_ordLess ordLess_imp_ordLeq False by blast  | 
707  | 
with assms Cinfinite_Cnotzero show ?thesis  | 
|
708  | 
by (force intro: regularCard_ordIso ordIso_symmetric[OF cprod_infinite1'])  | 
|
| 75624 | 709  | 
qed  | 
710  | 
||
| 54474 | 711  | 
lemma cprod_csum_cexp:  | 
712  | 
"r1 *c r2 \<le>o (r1 +c r2) ^c ctwo"  | 
|
| 76951 | 713  | 
unfolding cprod_def csum_def cexp_def ctwo_def Field_card_of  | 
| 54474 | 714  | 
proof -  | 
715  | 
let ?f = "\<lambda>(a, b). %x. if x then Inl a else Inr b"  | 
|
716  | 
have "inj_on ?f (Field r1 \<times> Field r2)" (is "inj_on _ ?LHS")  | 
|
717  | 
by (auto simp: inj_on_def fun_eq_iff split: bool.split)  | 
|
718  | 
moreover  | 
|
719  | 
have "?f ` ?LHS \<subseteq> Func (UNIV :: bool set) (Field r1 <+> Field r2)" (is "_ \<subseteq> ?RHS")  | 
|
720  | 
by (auto simp: Func_def)  | 
|
721  | 
ultimately show "|?LHS| \<le>o |?RHS|" using card_of_ordLeq by blast  | 
|
722  | 
qed  | 
|
723  | 
||
724  | 
lemma Cfinite_cprod_Cinfinite: "\<lbrakk>Cfinite r; Cinfinite s\<rbrakk> \<Longrightarrow> r *c s \<le>o s"  | 
|
| 76951 | 725  | 
by (intro cprod_cinfinite_bound)  | 
726  | 
(auto intro: ordLeq_refl ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite])  | 
|
| 54474 | 727  | 
|
728  | 
lemma cprod_cexp: "(r *c s) ^c t =o r ^c t *c s ^c t"  | 
|
729  | 
unfolding cprod_def cexp_def Field_card_of by (rule Func_Times_Range)  | 
|
730  | 
||
731  | 
lemma cprod_cexp_csum_cexp_Cinfinite:  | 
|
732  | 
assumes t: "Cinfinite t"  | 
|
733  | 
shows "(r *c s) ^c t \<le>o (r +c s) ^c t"  | 
|
734  | 
proof -  | 
|
735  | 
have "(r *c s) ^c t \<le>o ((r +c s) ^c ctwo) ^c t"  | 
|
736  | 
by (rule cexp_mono1[OF cprod_csum_cexp conjunct2[OF t]])  | 
|
737  | 
also have "((r +c s) ^c ctwo) ^c t =o (r +c s) ^c (ctwo *c t)"  | 
|
738  | 
by (rule cexp_cprod[OF Card_order_csum])  | 
|
739  | 
also have "(r +c s) ^c (ctwo *c t) =o (r +c s) ^c (t *c ctwo)"  | 
|
740  | 
by (rule cexp_cong2[OF cprod_com Card_order_csum Card_order_cprod])  | 
|
741  | 
also have "(r +c s) ^c (t *c ctwo) =o ((r +c s) ^c t) ^c ctwo"  | 
|
742  | 
by (rule ordIso_symmetric[OF cexp_cprod[OF Card_order_csum]])  | 
|
743  | 
also have "((r +c s) ^c t) ^c ctwo =o (r +c s) ^c t"  | 
|
744  | 
by (rule cexp_cprod_ordLeq[OF Card_order_csum t ctwo_Cnotzero ctwo_ordLeq_Cinfinite[OF t]])  | 
|
745  | 
finally show ?thesis .  | 
|
746  | 
qed  | 
|
747  | 
||
748  | 
lemma Cfinite_cexp_Cinfinite:  | 
|
749  | 
assumes s: "Cfinite s" and t: "Cinfinite t"  | 
|
750  | 
shows "s ^c t \<le>o ctwo ^c t"  | 
|
751  | 
proof (cases "s \<le>o ctwo")  | 
|
752  | 
case True thus ?thesis using t by (blast intro: cexp_mono1)  | 
|
753  | 
next  | 
|
754  | 
case False  | 
|
| 55811 | 755  | 
hence "ctwo \<le>o s" using ordLeq_total[of s ctwo] Card_order_ctwo s  | 
756  | 
by (auto intro: card_order_on_well_order_on)  | 
|
757  | 
hence "Cnotzero s" using Cnotzero_mono[OF ctwo_Cnotzero] s by blast  | 
|
758  | 
hence st: "Cnotzero (s *c t)" by (intro Cinfinite_Cnotzero[OF Cinfinite_cprod2]) (auto simp: t)  | 
|
| 54474 | 759  | 
have "s ^c t \<le>o (ctwo ^c s) ^c t"  | 
760  | 
using assms by (blast intro: cexp_mono1 ordLess_imp_ordLeq[OF ordLess_ctwo_cexp])  | 
|
761  | 
also have "(ctwo ^c s) ^c t =o ctwo ^c (s *c t)"  | 
|
762  | 
by (blast intro: Card_order_ctwo cexp_cprod)  | 
|
763  | 
also have "ctwo ^c (s *c t) \<le>o ctwo ^c t"  | 
|
764  | 
using assms st by (intro cexp_mono2_Cnotzero Cfinite_cprod_Cinfinite Card_order_ctwo)  | 
|
765  | 
finally show ?thesis .  | 
|
766  | 
qed  | 
|
767  | 
||
768  | 
lemma csum_Cfinite_cexp_Cinfinite:  | 
|
769  | 
assumes r: "Card_order r" and s: "Cfinite s" and t: "Cinfinite t"  | 
|
770  | 
shows "(r +c s) ^c t \<le>o (r +c ctwo) ^c t"  | 
|
771  | 
proof (cases "Cinfinite r")  | 
|
772  | 
case True  | 
|
773  | 
hence "r +c s =o r" by (intro csum_absorb1 ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite] s)  | 
|
774  | 
hence "(r +c s) ^c t =o r ^c t" using t by (blast intro: cexp_cong1)  | 
|
775  | 
also have "r ^c t \<le>o (r +c ctwo) ^c t" using t by (blast intro: cexp_mono1 ordLeq_csum1 r)  | 
|
776  | 
finally show ?thesis .  | 
|
777  | 
next  | 
|
778  | 
case False  | 
|
779  | 
with r have "Cfinite r" unfolding cinfinite_def cfinite_def by auto  | 
|
780  | 
hence "Cfinite (r +c s)" by (intro Cfinite_csum s)  | 
|
781  | 
hence "(r +c s) ^c t \<le>o ctwo ^c t" by (intro Cfinite_cexp_Cinfinite t)  | 
|
782  | 
also have "ctwo ^c t \<le>o (r +c ctwo) ^c t" using t  | 
|
783  | 
by (blast intro: cexp_mono1 ordLeq_csum2 Card_order_ctwo)  | 
|
784  | 
finally show ?thesis .  | 
|
785  | 
qed  | 
|
786  | 
||
787  | 
(* cardSuc *)  | 
|
788  | 
||
789  | 
lemma Cinfinite_cardSuc: "Cinfinite r \<Longrightarrow> Cinfinite (cardSuc r)"  | 
|
| 76951 | 790  | 
by (simp add: cinfinite_def cardSuc_Card_order cardSuc_finite)  | 
| 54474 | 791  | 
|
792  | 
lemma cardSuc_UNION_Cinfinite:  | 
|
| 69276 | 793  | 
assumes "Cinfinite r" "relChain (cardSuc r) As" "B \<le> (\<Union>i \<in> Field (cardSuc r). As i)" "|B| <=o r"  | 
| 67613 | 794  | 
shows "\<exists>i \<in> Field (cardSuc r). B \<le> As i"  | 
| 76951 | 795  | 
using cardSuc_UNION assms unfolding cinfinite_def by blast  | 
| 54474 | 796  | 
|
| 75624 | 797  | 
lemma Cinfinite_card_suc: "\<lbrakk> Cinfinite r ; card_order r \<rbrakk> \<Longrightarrow> Cinfinite (card_suc r)"  | 
798  | 
using Cinfinite_cong[OF cardSuc_ordIso_card_suc Cinfinite_cardSuc] .  | 
|
799  | 
||
| 
75625
 
0dd3ac5fdbaa
tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
 
traytel 
parents: 
75624 
diff
changeset
 | 
800  | 
lemma card_suc_least: "\<lbrakk>card_order r; Card_order s; r <o s\<rbrakk> \<Longrightarrow> card_suc r \<le>o s"  | 
| 
 
0dd3ac5fdbaa
tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
 
traytel 
parents: 
75624 
diff
changeset
 | 
801  | 
by (rule ordIso_ordLeq_trans[OF ordIso_symmetric[OF cardSuc_ordIso_card_suc]])  | 
| 
 
0dd3ac5fdbaa
tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
 
traytel 
parents: 
75624 
diff
changeset
 | 
802  | 
(auto intro!: cardSuc_least simp: card_order_on_Card_order)  | 
| 
 
0dd3ac5fdbaa
tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
 
traytel 
parents: 
75624 
diff
changeset
 | 
803  | 
|
| 75624 | 804  | 
lemma regularCard_cardSuc: "Cinfinite k \<Longrightarrow> regularCard (cardSuc k)"  | 
805  | 
by (rule infinite_cardSuc_regularCard) (auto simp: cinfinite_def)  | 
|
806  | 
||
| 
75625
 
0dd3ac5fdbaa
tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
 
traytel 
parents: 
75624 
diff
changeset
 | 
807  | 
lemma regularCard_card_suc: "card_order r \<Longrightarrow> Cinfinite r \<Longrightarrow> regularCard (card_suc r)"  | 
| 75624 | 808  | 
using cardSuc_ordIso_card_suc Cinfinite_cardSuc regularCard_cardSuc regularCard_ordIso  | 
809  | 
by blast  | 
|
810  | 
||
| 54474 | 811  | 
end  |