blanchet@55056: (* Title: HOL/BNF_Cardinal_Arithmetic.thy blanchet@54474: Author: Dmitriy Traytel, TU Muenchen blanchet@54474: Copyright 2012 blanchet@54474: blanchet@55059: Cardinal arithmetic as needed by bounded natural functors. blanchet@54474: *) blanchet@54474: wenzelm@60758: section \Cardinal Arithmetic as Needed by Bounded Natural Functors\ blanchet@54474: blanchet@55056: theory BNF_Cardinal_Arithmetic blanchet@55056: imports BNF_Cardinal_Order_Relation blanchet@54474: begin blanchet@54474: blanchet@54474: lemma dir_image: "\\x y. (f x = f y) = (x = y); Card_order r\ \ r =o dir_image r f" blanchet@54474: by (rule dir_image_ordIso) (auto simp add: inj_on_def card_order_on_def) blanchet@54474: blanchet@54474: lemma card_order_dir_image: blanchet@54474: assumes bij: "bij f" and co: "card_order r" blanchet@54474: shows "card_order (dir_image r f)" blanchet@54474: proof - blanchet@54474: from assms have "Field (dir_image r f) = UNIV" blanchet@54474: using card_order_on_Card_order[of UNIV r] unfolding bij_def dir_image_Field by auto blanchet@54474: moreover from bij have "\x y. (f x = f y) = (x = y)" unfolding bij_def inj_on_def by auto blanchet@54474: with co have "Card_order (dir_image r f)" blanchet@54474: using card_order_on_Card_order[of UNIV r] Card_order_ordIso2[OF _ dir_image] by blast blanchet@54474: ultimately show ?thesis by auto blanchet@54474: qed blanchet@54474: blanchet@54474: lemma ordIso_refl: "Card_order r \ r =o r" blanchet@54474: by (rule card_order_on_ordIso) blanchet@54474: blanchet@54474: lemma ordLeq_refl: "Card_order r \ r \o r" blanchet@54474: by (rule ordIso_imp_ordLeq, rule card_order_on_ordIso) blanchet@54474: blanchet@54474: lemma card_of_ordIso_subst: "A = B \ |A| =o |B|" blanchet@54474: by (simp only: ordIso_refl card_of_Card_order) blanchet@54474: blanchet@54474: lemma Field_card_order: "card_order r \ Field r = UNIV" blanchet@54474: using card_order_on_Card_order[of UNIV r] by simp blanchet@54474: blanchet@54474: wenzelm@60758: subsection \Zero\ blanchet@54474: blanchet@54474: definition czero where blanchet@54474: "czero = card_of {}" blanchet@54474: blanchet@54474: lemma czero_ordIso: blanchet@54474: "czero =o czero" blanchet@54474: using card_of_empty_ordIso by (simp add: czero_def) blanchet@54474: blanchet@54474: lemma card_of_ordIso_czero_iff_empty: blanchet@54474: "|A| =o (czero :: 'b rel) \ A = ({} :: 'a set)" blanchet@54474: unfolding czero_def by (rule iffI[OF card_of_empty2]) (auto simp: card_of_refl card_of_empty_ordIso) blanchet@54474: blanchet@54474: (* A "not czero" Cardinal predicate *) blanchet@54474: abbreviation Cnotzero where blanchet@54474: "Cnotzero (r :: 'a rel) \ \(r =o (czero :: 'a rel)) \ Card_order r" blanchet@54474: blanchet@54474: (*helper*) blanchet@54474: lemma Cnotzero_imp_not_empty: "Cnotzero r \ Field r \ {}" traytel@55811: unfolding Card_order_iff_ordIso_card_of czero_def by force blanchet@54474: blanchet@54474: lemma czeroI: blanchet@54474: "\Card_order r; Field r = {}\ \ r =o czero" blanchet@54474: using Cnotzero_imp_not_empty ordIso_transitive[OF _ czero_ordIso] by blast blanchet@54474: blanchet@54474: lemma czeroE: blanchet@54474: "r =o czero \ Field r = {}" blanchet@54474: unfolding czero_def blanchet@54474: by (drule card_of_cong) (simp only: Field_card_of card_of_empty2) blanchet@54474: blanchet@54474: lemma Cnotzero_mono: blanchet@54474: "\Cnotzero r; Card_order q; r \o q\ \ Cnotzero q" blanchet@54474: apply (rule ccontr) blanchet@54474: apply auto blanchet@54474: apply (drule czeroE) blanchet@54474: apply (erule notE) blanchet@54474: apply (erule czeroI) blanchet@54474: apply (drule card_of_mono2) blanchet@54474: apply (simp only: card_of_empty3) blanchet@54474: done blanchet@54474: wenzelm@60758: subsection \(In)finite cardinals\ blanchet@54474: blanchet@54474: definition cinfinite where traytel@54578: "cinfinite r = (\ finite (Field r))" blanchet@54474: blanchet@54474: abbreviation Cinfinite where blanchet@54474: "Cinfinite r \ cinfinite r \ Card_order r" blanchet@54474: blanchet@54474: definition cfinite where blanchet@54474: "cfinite r = finite (Field r)" blanchet@54474: blanchet@54474: abbreviation Cfinite where blanchet@54474: "Cfinite r \ cfinite r \ Card_order r" blanchet@54474: blanchet@54474: lemma Cfinite_ordLess_Cinfinite: "\Cfinite r; Cinfinite s\ \ r o r" blanchet@54474: proof - traytel@55811: from inf have "natLeq \o |Field r|" unfolding cinfinite_def traytel@55811: using infinite_iff_natLeq_ordLeq by blast blanchet@54474: also from inf have "|Field r| =o r" by (simp add: card_of_unique ordIso_symmetric) blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cinfinite_not_czero: "cinfinite r \ \ (r =o (czero :: 'a rel))" traytel@55811: unfolding cinfinite_def by (cases "Field r = {}") (auto dest: czeroE) blanchet@54474: blanchet@54474: lemma Cinfinite_Cnotzero: "Cinfinite r \ Cnotzero r" traytel@55811: by (rule conjI[OF cinfinite_not_czero]) simp_all blanchet@54474: blanchet@54474: lemma Cinfinite_cong: "\r1 =o r2; Cinfinite r1\ \ Cinfinite r2" traytel@55811: using Card_order_ordIso2[of r1 r2] unfolding cinfinite_def ordIso_iff_ordLeq traytel@55811: by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2]) blanchet@54474: blanchet@54474: lemma cinfinite_mono: "\r1 \o r2; cinfinite r1\ \ cinfinite r2" traytel@55811: unfolding cinfinite_def by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2]) blanchet@54474: blanchet@54474: wenzelm@60758: subsection \Binary sum\ blanchet@54474: blanchet@54474: definition csum (infixr "+c" 65) where blanchet@54474: "r1 +c r2 \ |Field r1 <+> Field r2|" blanchet@54474: blanchet@54474: lemma Field_csum: "Field (r +c s) = Inl ` Field r \ Inr ` Field s" blanchet@54474: unfolding csum_def Field_card_of by auto blanchet@54474: blanchet@54474: lemma Card_order_csum: blanchet@54474: "Card_order (r1 +c r2)" blanchet@54474: unfolding csum_def by (simp add: card_of_Card_order) blanchet@54474: blanchet@54474: lemma csum_Cnotzero1: blanchet@54474: "Cnotzero r1 \ Cnotzero (r1 +c r2)" traytel@55811: unfolding csum_def using Cnotzero_imp_not_empty[of r1] Plus_eq_empty_conv[of "Field r1" "Field r2"] traytel@55811: card_of_ordIso_czero_iff_empty[of "Field r1 <+> Field r2"] by (auto intro: card_of_Card_order) blanchet@54474: blanchet@54474: lemma card_order_csum: blanchet@54474: assumes "card_order r1" "card_order r2" blanchet@54474: shows "card_order (r1 +c r2)" blanchet@54474: proof - blanchet@54474: have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto blanchet@54474: thus ?thesis unfolding csum_def by (auto simp: card_of_card_order_on) blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cinfinite_csum: blanchet@54474: "cinfinite r1 \ cinfinite r2 \ cinfinite (r1 +c r2)" blanchet@54474: unfolding cinfinite_def csum_def by (auto simp: Field_card_of) blanchet@54474: blanchet@54474: lemma Cinfinite_csum1: blanchet@54474: "Cinfinite r1 \ Cinfinite (r1 +c r2)" traytel@55811: unfolding cinfinite_def csum_def by (rule conjI[OF _ card_of_Card_order]) (auto simp: Field_card_of) blanchet@54474: blanchet@54480: lemma Cinfinite_csum: blanchet@54480: "Cinfinite r1 \ Cinfinite r2 \ Cinfinite (r1 +c r2)" traytel@55811: unfolding cinfinite_def csum_def by (rule conjI[OF _ card_of_Card_order]) (auto simp: Field_card_of) blanchet@54480: blanchet@55851: lemma Cinfinite_csum_weak: blanchet@54480: "\Cinfinite r1; Cinfinite r2\ \ Cinfinite (r1 +c r2)" traytel@55811: by (erule Cinfinite_csum1) blanchet@54480: blanchet@54474: lemma csum_cong: "\p1 =o r1; p2 =o r2\ \ p1 +c p2 =o r1 +c r2" blanchet@54474: by (simp only: csum_def ordIso_Plus_cong) blanchet@54474: blanchet@54474: lemma csum_cong1: "p1 =o r1 \ p1 +c q =o r1 +c q" blanchet@54474: by (simp only: csum_def ordIso_Plus_cong1) blanchet@54474: blanchet@54474: lemma csum_cong2: "p2 =o r2 \ q +c p2 =o q +c r2" blanchet@54474: by (simp only: csum_def ordIso_Plus_cong2) blanchet@54474: blanchet@54474: lemma csum_mono: "\p1 \o r1; p2 \o r2\ \ p1 +c p2 \o r1 +c r2" blanchet@54474: by (simp only: csum_def ordLeq_Plus_mono) blanchet@54474: blanchet@54474: lemma csum_mono1: "p1 \o r1 \ p1 +c q \o r1 +c q" blanchet@54474: by (simp only: csum_def ordLeq_Plus_mono1) blanchet@54474: blanchet@54474: lemma csum_mono2: "p2 \o r2 \ q +c p2 \o q +c r2" blanchet@54474: by (simp only: csum_def ordLeq_Plus_mono2) blanchet@54474: blanchet@54474: lemma ordLeq_csum1: "Card_order p1 \ p1 \o p1 +c p2" blanchet@54474: by (simp only: csum_def Card_order_Plus1) blanchet@54474: blanchet@54474: lemma ordLeq_csum2: "Card_order p2 \ p2 \o p1 +c p2" blanchet@54474: by (simp only: csum_def Card_order_Plus2) blanchet@54474: blanchet@54474: lemma csum_com: "p1 +c p2 =o p2 +c p1" blanchet@54474: by (simp only: csum_def card_of_Plus_commute) blanchet@54474: blanchet@54474: lemma csum_assoc: "(p1 +c p2) +c p3 =o p1 +c p2 +c p3" blanchet@54474: by (simp only: csum_def Field_card_of card_of_Plus_assoc) blanchet@54474: blanchet@54474: lemma Cfinite_csum: "\Cfinite r; Cfinite s\ \ Cfinite (r +c s)" blanchet@54474: unfolding cfinite_def csum_def Field_card_of using card_of_card_order_on by simp blanchet@54474: blanchet@54474: lemma csum_csum: "(r1 +c r2) +c (r3 +c r4) =o (r1 +c r3) +c (r2 +c r4)" blanchet@54474: proof - blanchet@54474: have "(r1 +c r2) +c (r3 +c r4) =o r1 +c r2 +c (r3 +c r4)" traytel@55811: by (rule csum_assoc) blanchet@54474: also have "r1 +c r2 +c (r3 +c r4) =o r1 +c (r2 +c r3) +c r4" traytel@55811: by (intro csum_assoc csum_cong2 ordIso_symmetric) blanchet@54474: also have "r1 +c (r2 +c r3) +c r4 =o r1 +c (r3 +c r2) +c r4" traytel@55811: by (intro csum_com csum_cong1 csum_cong2) blanchet@54474: also have "r1 +c (r3 +c r2) +c r4 =o r1 +c r3 +c r2 +c r4" traytel@55811: by (intro csum_assoc csum_cong2 ordIso_symmetric) blanchet@54474: also have "r1 +c r3 +c r2 +c r4 =o (r1 +c r3) +c (r2 +c r4)" traytel@55811: by (intro csum_assoc ordIso_symmetric) blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: lemma Plus_csum: "|A <+> B| =o |A| +c |B|" blanchet@54474: by (simp only: csum_def Field_card_of card_of_refl) blanchet@54474: blanchet@54474: lemma Un_csum: "|A \ B| \o |A| +c |B|" blanchet@54474: using ordLeq_ordIso_trans[OF card_of_Un_Plus_ordLeq Plus_csum] by blast blanchet@54474: blanchet@54474: wenzelm@60758: subsection \One\ blanchet@54474: blanchet@54474: definition cone where blanchet@54474: "cone = card_of {()}" blanchet@54474: blanchet@54474: lemma Card_order_cone: "Card_order cone" blanchet@54474: unfolding cone_def by (rule card_of_Card_order) blanchet@54474: blanchet@54474: lemma Cfinite_cone: "Cfinite cone" blanchet@54474: unfolding cfinite_def by (simp add: Card_order_cone) blanchet@54474: blanchet@54474: lemma cone_not_czero: "\ (cone =o czero)" traytel@55811: unfolding czero_def cone_def ordIso_iff_ordLeq using card_of_empty3 empty_not_insert by blast blanchet@54474: blanchet@54474: lemma cone_ordLeq_Cnotzero: "Cnotzero r \ cone \o r" traytel@55811: unfolding cone_def by (rule Card_order_singl_ordLeq) (auto intro: czeroI) blanchet@54474: blanchet@54474: wenzelm@60758: subsection \Two\ blanchet@54474: blanchet@54474: definition ctwo where blanchet@54474: "ctwo = |UNIV :: bool set|" blanchet@54474: blanchet@54474: lemma Card_order_ctwo: "Card_order ctwo" blanchet@54474: unfolding ctwo_def by (rule card_of_Card_order) blanchet@54474: blanchet@54474: lemma ctwo_not_czero: "\ (ctwo =o czero)" blanchet@54474: using card_of_empty3[of "UNIV :: bool set"] ordIso_iff_ordLeq traytel@55811: unfolding czero_def ctwo_def using UNIV_not_empty by auto blanchet@54474: blanchet@54474: lemma ctwo_Cnotzero: "Cnotzero ctwo" blanchet@54474: by (simp add: ctwo_not_czero Card_order_ctwo) blanchet@54474: blanchet@54474: wenzelm@60758: subsection \Family sum\ blanchet@54474: blanchet@54474: definition Csum where blanchet@54474: "Csum r rs \ |SIGMA i : Field r. Field (rs i)|" blanchet@54474: blanchet@54474: (* Similar setup to the one for SIGMA from theory Big_Operators: *) blanchet@54474: syntax "_Csum" :: blanchet@54474: "pttrn => ('a * 'a) set => 'b * 'b set => (('a * 'b) * ('a * 'b)) set" blanchet@54474: ("(3CSUM _:_. _)" [0, 51, 10] 10) blanchet@54474: blanchet@54474: translations blanchet@54474: "CSUM i:r. rs" == "CONST Csum r (%i. rs)" blanchet@54474: blanchet@54474: lemma SIGMA_CSUM: "|SIGMA i : I. As i| = (CSUM i : |I|. |As i| )" blanchet@54474: by (auto simp: Csum_def Field_card_of) blanchet@54474: blanchet@54474: (* NB: Always, under the cardinal operator, blanchet@54474: operations on sets are reduced automatically to operations on cardinals. blanchet@54474: This should make cardinal reasoning more direct and natural. *) blanchet@54474: blanchet@54474: wenzelm@60758: subsection \Product\ blanchet@54474: blanchet@54474: definition cprod (infixr "*c" 80) where wenzelm@61943: "r1 *c r2 = |Field r1 \ Field r2|" blanchet@54474: blanchet@54474: lemma card_order_cprod: blanchet@54474: assumes "card_order r1" "card_order r2" blanchet@54474: shows "card_order (r1 *c r2)" blanchet@54474: proof - blanchet@54474: have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto blanchet@54474: thus ?thesis by (auto simp: cprod_def card_of_card_order_on) blanchet@54474: qed blanchet@54474: blanchet@54474: lemma Card_order_cprod: "Card_order (r1 *c r2)" blanchet@54474: by (simp only: cprod_def Field_card_of card_of_card_order_on) blanchet@54474: blanchet@54474: lemma cprod_mono1: "p1 \o r1 \ p1 *c q \o r1 *c q" blanchet@54474: by (simp only: cprod_def ordLeq_Times_mono1) blanchet@54474: blanchet@54474: lemma cprod_mono2: "p2 \o r2 \ q *c p2 \o q *c r2" blanchet@54474: by (simp only: cprod_def ordLeq_Times_mono2) blanchet@54474: blanchet@55851: lemma cprod_mono: "\p1 \o r1; p2 \o r2\ \ p1 *c p2 \o r1 *c r2" blanchet@55851: by (rule ordLeq_transitive[OF cprod_mono1 cprod_mono2]) blanchet@55851: blanchet@54474: lemma ordLeq_cprod2: "\Cnotzero p1; Card_order p2\ \ p2 \o p1 *c p2" traytel@55811: unfolding cprod_def by (rule Card_order_Times2) (auto intro: czeroI) blanchet@54474: blanchet@54474: lemma cinfinite_cprod: "\cinfinite r1; cinfinite r2\ \ cinfinite (r1 *c r2)" blanchet@54474: by (simp add: cinfinite_def cprod_def Field_card_of infinite_cartesian_product) blanchet@54474: blanchet@54474: lemma cinfinite_cprod2: "\Cnotzero r1; Cinfinite r2\ \ cinfinite (r1 *c r2)" traytel@55811: by (rule cinfinite_mono) (auto intro: ordLeq_cprod2) blanchet@54474: blanchet@54474: lemma Cinfinite_cprod2: "\Cnotzero r1; Cinfinite r2\ \ Cinfinite (r1 *c r2)" blanchet@54474: by (blast intro: cinfinite_cprod2 Card_order_cprod) blanchet@54474: blanchet@55851: lemma cprod_cong: "\p1 =o r1; p2 =o r2\ \ p1 *c p2 =o r1 *c r2" blanchet@55866: unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono) blanchet@55851: blanchet@55851: lemma cprod_cong1: "\p1 =o r1\ \ p1 *c p2 =o r1 *c p2" blanchet@55866: unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono1) blanchet@55851: blanchet@55851: lemma cprod_cong2: "p2 =o r2 \ q *c p2 =o q *c r2" blanchet@55866: unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono2) blanchet@55851: blanchet@54474: lemma cprod_com: "p1 *c p2 =o p2 *c p1" blanchet@54474: by (simp only: cprod_def card_of_Times_commute) blanchet@54474: blanchet@54474: lemma card_of_Csum_Times: blanchet@54474: "\i \ I. |A i| \o |B| \ (CSUM i : |I|. |A i| ) \o |I| *c |B|" traytel@56191: by (simp only: Csum_def cprod_def Field_card_of card_of_Sigma_mono1) blanchet@54474: blanchet@54474: lemma card_of_Csum_Times': blanchet@54474: assumes "Card_order r" "\i \ I. |A i| \o r" blanchet@54474: shows "(CSUM i : |I|. |A i| ) \o |I| *c r" blanchet@54474: proof - blanchet@54474: from assms(1) have *: "r =o |Field r|" by (simp add: card_of_unique) blanchet@54474: with assms(2) have "\i \ I. |A i| \o |Field r|" by (blast intro: ordLeq_ordIso_trans) blanchet@54474: hence "(CSUM i : |I|. |A i| ) \o |I| *c |Field r|" by (simp only: card_of_Csum_Times) blanchet@54474: also from * have "|I| *c |Field r| \o |I| *c r" blanchet@54474: by (simp only: Field_card_of card_of_refl cprod_def ordIso_imp_ordLeq) blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cprod_csum_distrib1: "r1 *c r2 +c r1 *c r3 =o r1 *c (r2 +c r3)" blanchet@54474: unfolding csum_def cprod_def by (simp add: Field_card_of card_of_Times_Plus_distrib ordIso_symmetric) blanchet@54474: blanchet@54474: lemma csum_absorb2': "\Card_order r2; r1 \o r2; cinfinite r1 \ cinfinite r2\ \ r1 +c r2 =o r2" traytel@55811: unfolding csum_def by (rule conjunct2[OF Card_order_Plus_infinite]) traytel@55811: (auto simp: cinfinite_def dest: cinfinite_mono) blanchet@54474: blanchet@54474: lemma csum_absorb1': blanchet@54474: assumes card: "Card_order r2" blanchet@54474: and r12: "r1 \o r2" and cr12: "cinfinite r1 \ cinfinite r2" blanchet@54474: shows "r2 +c r1 =o r2" blanchet@54474: by (rule ordIso_transitive, rule csum_com, rule csum_absorb2', (simp only: assms)+) blanchet@54474: blanchet@54474: lemma csum_absorb1: "\Cinfinite r2; r1 \o r2\ \ r2 +c r1 =o r2" blanchet@54474: by (rule csum_absorb1') auto blanchet@54474: blanchet@54474: wenzelm@60758: subsection \Exponentiation\ blanchet@54474: blanchet@54474: definition cexp (infixr "^c" 90) where blanchet@54474: "r1 ^c r2 \ |Func (Field r2) (Field r1)|" blanchet@54474: blanchet@54474: lemma Card_order_cexp: "Card_order (r1 ^c r2)" blanchet@54474: unfolding cexp_def by (rule card_of_Card_order) blanchet@54474: blanchet@54474: lemma cexp_mono': blanchet@54474: assumes 1: "p1 \o r1" and 2: "p2 \o r2" blanchet@54474: and n: "Field p2 = {} \ Field r2 = {}" blanchet@54474: shows "p1 ^c p2 \o r1 ^c r2" blanchet@54474: proof(cases "Field p1 = {}") blanchet@54474: case True traytel@55811: hence "Field p2 \ {} \ Func (Field p2) {} = {}" unfolding Func_is_emp by simp traytel@55811: with True have "|Field |Func (Field p2) (Field p1)|| \o cone" blanchet@54474: unfolding cone_def Field_card_of traytel@55811: by (cases "Field p2 = {}", auto intro: surj_imp_ordLeq simp: Func_empty) blanchet@54474: hence "|Func (Field p2) (Field p1)| \o cone" by (simp add: Field_card_of cexp_def) blanchet@54474: hence "p1 ^c p2 \o cone" unfolding cexp_def . blanchet@54474: thus ?thesis blanchet@54474: proof (cases "Field p2 = {}") blanchet@54474: case True blanchet@54474: with n have "Field r2 = {}" . noschinl@55604: hence "cone \o r1 ^c r2" unfolding cone_def cexp_def Func_def noschinl@55604: by (auto intro: card_of_ordLeqI[where f="\_ _. undefined"]) wenzelm@60758: thus ?thesis using \p1 ^c p2 \o cone\ ordLeq_transitive by auto blanchet@54474: next blanchet@54474: case False with True have "|Field (p1 ^c p2)| =o czero" blanchet@54474: unfolding card_of_ordIso_czero_iff_empty cexp_def Field_card_of Func_def by auto blanchet@54474: thus ?thesis unfolding cexp_def card_of_ordIso_czero_iff_empty Field_card_of blanchet@54474: by (simp add: card_of_empty) blanchet@54474: qed blanchet@54474: next blanchet@54474: case False blanchet@54474: have 1: "|Field p1| \o |Field r1|" and 2: "|Field p2| \o |Field r2|" blanchet@54474: using 1 2 by (auto simp: card_of_mono2) blanchet@54474: obtain f1 where f1: "f1 ` Field r1 = Field p1" blanchet@54474: using 1 unfolding card_of_ordLeq2[OF False, symmetric] by auto blanchet@54474: obtain f2 where f2: "inj_on f2 (Field p2)" "f2 ` Field p2 \ Field r2" blanchet@54474: using 2 unfolding card_of_ordLeq[symmetric] by blast blanchet@54474: have 0: "Func_map (Field p2) f1 f2 ` (Field (r1 ^c r2)) = Field (p1 ^c p2)" blanchet@54474: unfolding cexp_def Field_card_of using Func_map_surj[OF f1 f2 n, symmetric] . blanchet@54474: have 00: "Field (p1 ^c p2) \ {}" unfolding cexp_def Field_card_of Func_is_emp blanchet@54474: using False by simp blanchet@54474: show ?thesis blanchet@54474: using 0 card_of_ordLeq2[OF 00] unfolding cexp_def Field_card_of by blast blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cexp_mono: blanchet@54474: assumes 1: "p1 \o r1" and 2: "p2 \o r2" blanchet@54474: and n: "p2 =o czero \ r2 =o czero" and card: "Card_order p2" blanchet@54474: shows "p1 ^c p2 \o r1 ^c r2" traytel@55811: by (rule cexp_mono'[OF 1 2 czeroE[OF n[OF czeroI[OF card]]]]) blanchet@54474: blanchet@54474: lemma cexp_mono1: blanchet@54474: assumes 1: "p1 \o r1" and q: "Card_order q" blanchet@54474: shows "p1 ^c q \o r1 ^c q" blanchet@54474: using ordLeq_refl[OF q] by (rule cexp_mono[OF 1]) (auto simp: q) blanchet@54474: blanchet@54474: lemma cexp_mono2': blanchet@54474: assumes 2: "p2 \o r2" and q: "Card_order q" blanchet@54474: and n: "Field p2 = {} \ Field r2 = {}" blanchet@54474: shows "q ^c p2 \o q ^c r2" blanchet@54474: using ordLeq_refl[OF q] by (rule cexp_mono'[OF _ 2 n]) auto blanchet@54474: blanchet@54474: lemma cexp_mono2: blanchet@54474: assumes 2: "p2 \o r2" and q: "Card_order q" blanchet@54474: and n: "p2 =o czero \ r2 =o czero" and card: "Card_order p2" blanchet@54474: shows "q ^c p2 \o q ^c r2" blanchet@54474: using ordLeq_refl[OF q] by (rule cexp_mono[OF _ 2 n card]) auto blanchet@54474: blanchet@54474: lemma cexp_mono2_Cnotzero: blanchet@54474: assumes "p2 \o r2" "Card_order q" "Cnotzero p2" blanchet@54474: shows "q ^c p2 \o q ^c r2" traytel@55811: using assms(3) czeroI by (blast intro: cexp_mono2'[OF assms(1,2)]) blanchet@54474: blanchet@54474: lemma cexp_cong: blanchet@54474: assumes 1: "p1 =o r1" and 2: "p2 =o r2" blanchet@54474: and Cr: "Card_order r2" blanchet@54474: and Cp: "Card_order p2" blanchet@54474: shows "p1 ^c p2 =o r1 ^c r2" blanchet@54474: proof - blanchet@54474: obtain f where "bij_betw f (Field p2) (Field r2)" blanchet@54474: using 2 card_of_ordIso[of "Field p2" "Field r2"] card_of_cong by auto blanchet@54474: hence 0: "Field p2 = {} \ Field r2 = {}" unfolding bij_betw_def by auto blanchet@54474: have r: "p2 =o czero \ r2 =o czero" blanchet@54474: and p: "r2 =o czero \ p2 =o czero" blanchet@54474: using 0 Cr Cp czeroE czeroI by auto blanchet@54474: show ?thesis using 0 1 2 unfolding ordIso_iff_ordLeq traytel@55811: using r p cexp_mono[OF _ _ _ Cp] cexp_mono[OF _ _ _ Cr] by blast blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cexp_cong1: blanchet@54474: assumes 1: "p1 =o r1" and q: "Card_order q" blanchet@54474: shows "p1 ^c q =o r1 ^c q" blanchet@54474: by (rule cexp_cong[OF 1 _ q q]) (rule ordIso_refl[OF q]) blanchet@54474: blanchet@54474: lemma cexp_cong2: blanchet@54474: assumes 2: "p2 =o r2" and q: "Card_order q" and p: "Card_order p2" blanchet@54474: shows "q ^c p2 =o q ^c r2" blanchet@54474: by (rule cexp_cong[OF _ 2]) (auto simp only: ordIso_refl Card_order_ordIso2[OF p 2] q p) blanchet@54474: blanchet@54474: lemma cexp_cone: blanchet@54474: assumes "Card_order r" blanchet@54474: shows "r ^c cone =o r" blanchet@54474: proof - blanchet@54474: have "r ^c cone =o |Field r|" blanchet@54474: unfolding cexp_def cone_def Field_card_of Func_empty blanchet@54474: card_of_ordIso[symmetric] bij_betw_def Func_def inj_on_def image_def blanchet@54474: by (rule exI[of _ "\f. f ()"]) auto blanchet@54474: also have "|Field r| =o r" by (rule card_of_Field_ordIso[OF assms]) blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cexp_cprod: blanchet@54474: assumes r1: "Card_order r1" blanchet@54474: shows "(r1 ^c r2) ^c r3 =o r1 ^c (r2 *c r3)" (is "?L =o ?R") blanchet@54474: proof - blanchet@54474: have "?L =o r1 ^c (r3 *c r2)" blanchet@54474: unfolding cprod_def cexp_def Field_card_of blanchet@54474: using card_of_Func_Times by(rule ordIso_symmetric) blanchet@54474: also have "r1 ^c (r3 *c r2) =o ?R" blanchet@54474: apply(rule cexp_cong2) using cprod_com r1 by (auto simp: Card_order_cprod) blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cprod_infinite1': "\Cinfinite r; Cnotzero p; p \o r\ \ r *c p =o r" blanchet@54474: unfolding cinfinite_def cprod_def blanchet@54474: by (rule Card_order_Times_infinite[THEN conjunct1]) (blast intro: czeroI)+ blanchet@54474: blanchet@55851: lemma cprod_infinite: "Cinfinite r \ r *c r =o r" blanchet@55851: using cprod_infinite1' Cinfinite_Cnotzero ordLeq_refl by blast blanchet@55851: blanchet@54474: lemma cexp_cprod_ordLeq: blanchet@54474: assumes r1: "Card_order r1" and r2: "Cinfinite r2" blanchet@54474: and r3: "Cnotzero r3" "r3 \o r2" blanchet@54474: shows "(r1 ^c r2) ^c r3 =o r1 ^c r2" (is "?L =o ?R") blanchet@54474: proof- blanchet@54474: have "?L =o r1 ^c (r2 *c r3)" using cexp_cprod[OF r1] . blanchet@54474: also have "r1 ^c (r2 *c r3) =o ?R" blanchet@54474: apply(rule cexp_cong2) blanchet@54474: apply(rule cprod_infinite1'[OF r2 r3]) using r1 r2 by (fastforce simp: Card_order_cprod)+ blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: lemma Cnotzero_UNIV: "Cnotzero |UNIV|" blanchet@54474: by (auto simp: card_of_Card_order card_of_ordIso_czero_iff_empty) blanchet@54474: blanchet@54474: lemma ordLess_ctwo_cexp: blanchet@54474: assumes "Card_order r" blanchet@54474: shows "r o q ^c r" blanchet@54474: proof (cases "q =o (czero :: 'a rel)") blanchet@54474: case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans) blanchet@54474: next blanchet@54474: case False blanchet@54474: thus ?thesis blanchet@54474: apply - blanchet@54474: apply (rule ordIso_ordLeq_trans) blanchet@54474: apply (rule ordIso_symmetric) blanchet@54474: apply (rule cexp_cone) blanchet@54474: apply (rule assms(2)) blanchet@54474: apply (rule cexp_mono2) blanchet@54474: apply (rule cone_ordLeq_Cnotzero) blanchet@54474: apply (rule assms(1)) blanchet@54474: apply (rule assms(2)) blanchet@54474: apply (rule notE) blanchet@54474: apply (rule cone_not_czero) blanchet@54474: apply assumption blanchet@54474: apply (rule Card_order_cone) blanchet@54474: done blanchet@54474: qed blanchet@54474: blanchet@54474: lemma ordLeq_cexp2: blanchet@54474: assumes "ctwo \o q" "Card_order r" blanchet@54474: shows "r \o q ^c r" blanchet@54474: proof (cases "r =o (czero :: 'a rel)") blanchet@54474: case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans) blanchet@54474: next blanchet@54474: case False thus ?thesis blanchet@54474: apply - blanchet@54474: apply (rule ordLess_imp_ordLeq) blanchet@54474: apply (rule ordLess_ordLeq_trans) blanchet@54474: apply (rule ordLess_ctwo_cexp) blanchet@54474: apply (rule assms(2)) blanchet@54474: apply (rule cexp_mono1) blanchet@54474: apply (rule assms(1)) blanchet@54474: apply (rule assms(2)) blanchet@54474: done blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cinfinite_cexp: "\ctwo \o q; Cinfinite r\ \ cinfinite (q ^c r)" traytel@55811: by (rule cinfinite_mono[OF ordLeq_cexp2]) simp_all blanchet@54474: blanchet@54474: lemma Cinfinite_cexp: blanchet@54474: "\ctwo \o q; Cinfinite r\ \ Cinfinite (q ^c r)" blanchet@54474: by (simp add: cinfinite_cexp Card_order_cexp) blanchet@54474: blanchet@54474: lemma ctwo_ordLess_natLeq: "ctwo ctwo o r" blanchet@54474: by (rule ordLess_imp_ordLeq[OF ctwo_ordLess_Cinfinite[OF assms]]) blanchet@54474: blanchet@54474: lemma Un_Cinfinite_bound: "\|A| \o r; |B| \o r; Cinfinite r\ \ |A \ B| \o r" blanchet@54474: by (auto simp add: cinfinite_def card_of_Un_ordLeq_infinite_Field) blanchet@54474: blanchet@54474: lemma UNION_Cinfinite_bound: "\|I| \o r; \i \ I. |A i| \o r; Cinfinite r\ \ |\i \ I. A i| \o r" blanchet@54474: by (auto simp add: card_of_UNION_ordLeq_infinite_Field cinfinite_def) blanchet@54474: blanchet@54474: lemma csum_cinfinite_bound: blanchet@54474: assumes "p \o r" "q \o r" "Card_order p" "Card_order q" "Cinfinite r" blanchet@54474: shows "p +c q \o r" blanchet@54474: proof - blanchet@54474: from assms(1-4) have "|Field p| \o r" "|Field q| \o r" blanchet@54474: unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+ blanchet@54474: with assms show ?thesis unfolding cinfinite_def csum_def blanchet@54474: by (blast intro: card_of_Plus_ordLeq_infinite_Field) blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cprod_cinfinite_bound: blanchet@54474: assumes "p \o r" "q \o r" "Card_order p" "Card_order q" "Cinfinite r" blanchet@54474: shows "p *c q \o r" blanchet@54474: proof - blanchet@54474: from assms(1-4) have "|Field p| \o r" "|Field q| \o r" blanchet@54474: unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+ blanchet@54474: with assms show ?thesis unfolding cinfinite_def cprod_def blanchet@54474: by (blast intro: card_of_Times_ordLeq_infinite_Field) blanchet@54474: qed blanchet@54474: blanchet@54474: lemma cprod_csum_cexp: blanchet@54474: "r1 *c r2 \o (r1 +c r2) ^c ctwo" blanchet@54474: unfolding cprod_def csum_def cexp_def ctwo_def Field_card_of blanchet@54474: proof - blanchet@54474: let ?f = "\(a, b). %x. if x then Inl a else Inr b" blanchet@54474: have "inj_on ?f (Field r1 \ Field r2)" (is "inj_on _ ?LHS") blanchet@54474: by (auto simp: inj_on_def fun_eq_iff split: bool.split) blanchet@54474: moreover blanchet@54474: have "?f ` ?LHS \ Func (UNIV :: bool set) (Field r1 <+> Field r2)" (is "_ \ ?RHS") blanchet@54474: by (auto simp: Func_def) blanchet@54474: ultimately show "|?LHS| \o |?RHS|" using card_of_ordLeq by blast blanchet@54474: qed blanchet@54474: blanchet@54474: lemma Cfinite_cprod_Cinfinite: "\Cfinite r; Cinfinite s\ \ r *c s \o s" blanchet@54474: by (intro cprod_cinfinite_bound) blanchet@54474: (auto intro: ordLeq_refl ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite]) blanchet@54474: blanchet@54474: lemma cprod_cexp: "(r *c s) ^c t =o r ^c t *c s ^c t" blanchet@54474: unfolding cprod_def cexp_def Field_card_of by (rule Func_Times_Range) blanchet@54474: blanchet@54474: lemma cprod_cexp_csum_cexp_Cinfinite: blanchet@54474: assumes t: "Cinfinite t" blanchet@54474: shows "(r *c s) ^c t \o (r +c s) ^c t" blanchet@54474: proof - blanchet@54474: have "(r *c s) ^c t \o ((r +c s) ^c ctwo) ^c t" blanchet@54474: by (rule cexp_mono1[OF cprod_csum_cexp conjunct2[OF t]]) blanchet@54474: also have "((r +c s) ^c ctwo) ^c t =o (r +c s) ^c (ctwo *c t)" blanchet@54474: by (rule cexp_cprod[OF Card_order_csum]) blanchet@54474: also have "(r +c s) ^c (ctwo *c t) =o (r +c s) ^c (t *c ctwo)" blanchet@54474: by (rule cexp_cong2[OF cprod_com Card_order_csum Card_order_cprod]) blanchet@54474: also have "(r +c s) ^c (t *c ctwo) =o ((r +c s) ^c t) ^c ctwo" blanchet@54474: by (rule ordIso_symmetric[OF cexp_cprod[OF Card_order_csum]]) blanchet@54474: also have "((r +c s) ^c t) ^c ctwo =o (r +c s) ^c t" blanchet@54474: by (rule cexp_cprod_ordLeq[OF Card_order_csum t ctwo_Cnotzero ctwo_ordLeq_Cinfinite[OF t]]) blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: lemma Cfinite_cexp_Cinfinite: blanchet@54474: assumes s: "Cfinite s" and t: "Cinfinite t" blanchet@54474: shows "s ^c t \o ctwo ^c t" blanchet@54474: proof (cases "s \o ctwo") blanchet@54474: case True thus ?thesis using t by (blast intro: cexp_mono1) blanchet@54474: next blanchet@54474: case False traytel@55811: hence "ctwo \o s" using ordLeq_total[of s ctwo] Card_order_ctwo s traytel@55811: by (auto intro: card_order_on_well_order_on) traytel@55811: hence "Cnotzero s" using Cnotzero_mono[OF ctwo_Cnotzero] s by blast traytel@55811: hence st: "Cnotzero (s *c t)" by (intro Cinfinite_Cnotzero[OF Cinfinite_cprod2]) (auto simp: t) blanchet@54474: have "s ^c t \o (ctwo ^c s) ^c t" blanchet@54474: using assms by (blast intro: cexp_mono1 ordLess_imp_ordLeq[OF ordLess_ctwo_cexp]) blanchet@54474: also have "(ctwo ^c s) ^c t =o ctwo ^c (s *c t)" blanchet@54474: by (blast intro: Card_order_ctwo cexp_cprod) blanchet@54474: also have "ctwo ^c (s *c t) \o ctwo ^c t" blanchet@54474: using assms st by (intro cexp_mono2_Cnotzero Cfinite_cprod_Cinfinite Card_order_ctwo) blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: lemma csum_Cfinite_cexp_Cinfinite: blanchet@54474: assumes r: "Card_order r" and s: "Cfinite s" and t: "Cinfinite t" blanchet@54474: shows "(r +c s) ^c t \o (r +c ctwo) ^c t" blanchet@54474: proof (cases "Cinfinite r") blanchet@54474: case True blanchet@54474: hence "r +c s =o r" by (intro csum_absorb1 ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite] s) blanchet@54474: hence "(r +c s) ^c t =o r ^c t" using t by (blast intro: cexp_cong1) blanchet@54474: also have "r ^c t \o (r +c ctwo) ^c t" using t by (blast intro: cexp_mono1 ordLeq_csum1 r) blanchet@54474: finally show ?thesis . blanchet@54474: next blanchet@54474: case False blanchet@54474: with r have "Cfinite r" unfolding cinfinite_def cfinite_def by auto blanchet@54474: hence "Cfinite (r +c s)" by (intro Cfinite_csum s) blanchet@54474: hence "(r +c s) ^c t \o ctwo ^c t" by (intro Cfinite_cexp_Cinfinite t) blanchet@54474: also have "ctwo ^c t \o (r +c ctwo) ^c t" using t blanchet@54474: by (blast intro: cexp_mono1 ordLeq_csum2 Card_order_ctwo) blanchet@54474: finally show ?thesis . blanchet@54474: qed blanchet@54474: blanchet@54474: (* cardSuc *) blanchet@54474: blanchet@54474: lemma Cinfinite_cardSuc: "Cinfinite r \ Cinfinite (cardSuc r)" blanchet@54474: by (simp add: cinfinite_def cardSuc_Card_order cardSuc_finite) blanchet@54474: blanchet@54474: lemma cardSuc_UNION_Cinfinite: blanchet@54474: assumes "Cinfinite r" "relChain (cardSuc r) As" "B \ (UN i : Field (cardSuc r). As i)" "|B| <=o r" blanchet@54474: shows "EX i : Field (cardSuc r). B \ As i" blanchet@54474: using cardSuc_UNION assms unfolding cinfinite_def by blast blanchet@54474: blanchet@54474: end