blanchet@49310: (* Title: HOL/Cardinals/Cardinal_Arithmetic.thy blanchet@48975: Author: Dmitriy Traytel, TU Muenchen blanchet@48975: Copyright 2012 blanchet@48975: blanchet@48975: Cardinal arithmetic. blanchet@48975: *) blanchet@48975: wenzelm@58889: section {* Cardinal Arithmetic *} blanchet@48975: blanchet@48975: theory Cardinal_Arithmetic blanchet@55056: imports BNF_Cardinal_Arithmetic Cardinal_Order_Relation blanchet@48975: begin blanchet@48975: blanchet@48975: subsection {* Binary sum *} blanchet@48975: blanchet@48975: lemma csum_Cnotzero2: blanchet@48975: "Cnotzero r2 \ Cnotzero (r1 +c r2)" blanchet@48975: unfolding csum_def blanchet@48975: by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE) blanchet@48975: blanchet@48975: lemma single_cone: blanchet@48975: "|{x}| =o cone" blanchet@48975: proof - blanchet@48975: let ?f = "\x. ()" blanchet@48975: have "bij_betw ?f {x} {()}" unfolding bij_betw_def by auto blanchet@48975: thus ?thesis unfolding cone_def using card_of_ordIso by blast blanchet@48975: qed blanchet@48975: blanchet@48975: lemma cone_Cnotzero: "Cnotzero cone" blanchet@48975: by (simp add: cone_not_czero Card_order_cone) blanchet@48975: blanchet@48975: lemma cone_ordLeq_ctwo: "cone \o ctwo" blanchet@48975: unfolding cone_def ctwo_def card_of_ordLeq[symmetric] by auto blanchet@48975: traytel@54980: lemma csum_czero1: "Card_order r \ r +c czero =o r" traytel@54980: unfolding czero_def csum_def Field_card_of traytel@54980: by (rule ordIso_transitive[OF ordIso_symmetric[OF card_of_Plus_empty1] card_of_Field_ordIso]) traytel@54980: traytel@54980: lemma csum_czero2: "Card_order r \ czero +c r =o r" traytel@54980: unfolding czero_def csum_def Field_card_of traytel@54980: by (rule ordIso_transitive[OF ordIso_symmetric[OF card_of_Plus_empty2] card_of_Field_ordIso]) traytel@54980: blanchet@48975: blanchet@48975: subsection {* Product *} blanchet@48975: blanchet@48975: lemma Times_cprod: "|A \ B| =o |A| *c |B|" blanchet@48975: by (simp only: cprod_def Field_card_of card_of_refl) blanchet@48975: traytel@54980: lemma card_of_Times_singleton: traytel@54980: fixes A :: "'a set" traytel@54980: shows "|A \ {x}| =o |A|" traytel@54980: proof - traytel@54980: def f \ "\(a::'a,b::'b). (a)" wenzelm@61943: have "A \ f ` (A \ {x})" unfolding f_def by (auto simp: image_iff) wenzelm@61943: hence "bij_betw f (A \ {x}) A" unfolding bij_betw_def inj_on_def f_def by fastforce traytel@54980: thus ?thesis using card_of_ordIso by blast traytel@54980: qed traytel@54980: traytel@54980: lemma cprod_assoc: "(r *c s) *c t =o r *c s *c t" traytel@54980: unfolding cprod_def Field_card_of by (rule card_of_Times_assoc) traytel@54980: traytel@54980: lemma cprod_czero: "r *c czero =o czero" traytel@54980: unfolding cprod_def czero_def Field_card_of by (simp add: card_of_empty_ordIso) traytel@54980: traytel@54980: lemma cprod_cone: "Card_order r \ r *c cone =o r" traytel@54980: unfolding cprod_def cone_def Field_card_of traytel@54980: by (drule card_of_Field_ordIso) (erule ordIso_transitive[OF card_of_Times_singleton]) traytel@54980: blanchet@48975: blanchet@48975: lemma ordLeq_cprod1: "\Card_order p1; Cnotzero p2\ \ p1 \o p1 *c p2" blanchet@48975: unfolding cprod_def by (metis Card_order_Times1 czeroI) blanchet@48975: blanchet@48975: blanchet@48975: subsection {* Exponentiation *} blanchet@48975: blanchet@48975: lemma cexp_czero: "r ^c czero =o cone" blanchet@48975: unfolding cexp_def czero_def Field_card_of Func_empty by (rule single_cone) blanchet@48975: blanchet@48975: lemma Pow_cexp_ctwo: blanchet@48975: "|Pow A| =o ctwo ^c |A|" blanchet@48975: unfolding ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func) blanchet@48975: blanchet@48975: lemma Cnotzero_cexp: blanchet@48975: assumes "Cnotzero q" "Card_order r" blanchet@48975: shows "Cnotzero (q ^c r)" blanchet@48975: proof (cases "r =o czero") blanchet@48975: case False thus ?thesis blanchet@48975: apply - blanchet@48975: apply (rule Cnotzero_mono) blanchet@48975: apply (rule assms(1)) blanchet@48975: apply (rule Card_order_cexp) blanchet@48975: apply (rule ordLeq_cexp1) blanchet@48975: apply (rule conjI) blanchet@48975: apply (rule notI) blanchet@48975: apply (erule notE) blanchet@48975: apply (rule ordIso_transitive) blanchet@48975: apply assumption blanchet@48975: apply (rule czero_ordIso) blanchet@48975: apply (rule assms(2)) blanchet@48975: apply (rule conjunct2) blanchet@48975: apply (rule assms(1)) blanchet@48975: done blanchet@48975: next blanchet@48975: case True thus ?thesis blanchet@48975: apply - blanchet@48975: apply (rule Cnotzero_mono) blanchet@48975: apply (rule cone_Cnotzero) blanchet@48975: apply (rule Card_order_cexp) blanchet@48975: apply (rule ordIso_imp_ordLeq) blanchet@48975: apply (rule ordIso_symmetric) blanchet@48975: apply (rule ordIso_transitive) blanchet@48975: apply (rule cexp_cong2) blanchet@48975: apply assumption blanchet@48975: apply (rule conjunct2) blanchet@48975: apply (rule assms(1)) blanchet@48975: apply (rule assms(2)) blanchet@48975: apply (rule cexp_czero) blanchet@48975: done blanchet@48975: qed blanchet@48975: blanchet@48975: lemma Cinfinite_ctwo_cexp: blanchet@48975: "Cinfinite r \ Cinfinite (ctwo ^c r)" blanchet@48975: unfolding ctwo_def cexp_def cinfinite_def Field_card_of blanchet@54475: by (rule conjI, rule infinite_Func, auto) blanchet@48975: blanchet@48975: lemma cone_ordLeq_iff_Field: blanchet@48975: assumes "cone \o r" blanchet@48975: shows "Field r \ {}" blanchet@48975: proof (rule ccontr) blanchet@48975: assume "\ Field r \ {}" blanchet@48975: hence "Field r = {}" by simp blanchet@48975: thus False using card_of_empty3 blanchet@48975: card_of_mono2[OF assms] Cnotzero_imp_not_empty[OF cone_Cnotzero] by auto blanchet@48975: qed blanchet@48975: blanchet@48975: lemma cone_ordLeq_cexp: "cone \o r1 \ cone \o r1 ^c r2" traytel@54980: by (simp add: cexp_def cone_def Func_non_emp cone_ordLeq_iff_Field) blanchet@48975: blanchet@48975: lemma Card_order_czero: "Card_order czero" blanchet@48975: by (simp only: card_of_Card_order czero_def) blanchet@48975: blanchet@48975: lemma cexp_mono2'': blanchet@48975: assumes 2: "p2 \o r2" blanchet@48975: and n1: "Cnotzero q" blanchet@48975: and n2: "Card_order p2" blanchet@48975: shows "q ^c p2 \o q ^c r2" blanchet@48975: proof (cases "p2 =o (czero :: 'a rel)") blanchet@48975: case True blanchet@48975: hence "q ^c p2 =o q ^c (czero :: 'a rel)" using n1 n2 cexp_cong2 Card_order_czero by blast blanchet@48975: also have "q ^c (czero :: 'a rel) =o cone" using cexp_czero by blast blanchet@48975: also have "cone \o q ^c r2" using cone_ordLeq_cexp cone_ordLeq_Cnotzero n1 by blast blanchet@48975: finally show ?thesis . blanchet@48975: next blanchet@48975: case False thus ?thesis using assms cexp_mono2' czeroI by metis blanchet@48975: qed blanchet@48975: blanchet@48975: lemma csum_cexp: "\Cinfinite r1; Cinfinite r2; Card_order q; ctwo \o q\ \ blanchet@48975: q ^c r1 +c q ^c r2 \o q ^c (r1 +c r2)" blanchet@48975: apply (rule csum_cinfinite_bound) blanchet@48975: apply (rule cexp_mono2) blanchet@48975: apply (rule ordLeq_csum1) blanchet@48975: apply (erule conjunct2) blanchet@48975: apply assumption blanchet@48975: apply (rule notE) blanchet@48975: apply (rule cinfinite_not_czero[of r1]) blanchet@48975: apply (erule conjunct1) blanchet@48975: apply assumption blanchet@48975: apply (erule conjunct2) blanchet@48975: apply (rule cexp_mono2) blanchet@48975: apply (rule ordLeq_csum2) blanchet@48975: apply (erule conjunct2) blanchet@48975: apply assumption blanchet@48975: apply (rule notE) blanchet@48975: apply (rule cinfinite_not_czero[of r2]) blanchet@48975: apply (erule conjunct1) blanchet@48975: apply assumption blanchet@48975: apply (erule conjunct2) blanchet@48975: apply (rule Card_order_cexp) blanchet@48975: apply (rule Card_order_cexp) blanchet@48975: apply (rule Cinfinite_cexp) blanchet@48975: apply assumption blanchet@48975: apply (rule Cinfinite_csum) blanchet@48975: by (rule disjI1) blanchet@48975: blanchet@48975: lemma csum_cexp': "\Cinfinite r; Card_order q; ctwo \o q\ \ q +c r \o q ^c r" blanchet@48975: apply (rule csum_cinfinite_bound) blanchet@48975: apply (metis Cinfinite_Cnotzero ordLeq_cexp1) blanchet@48975: apply (metis ordLeq_cexp2) blanchet@48975: apply blast+ blanchet@48975: by (metis Cinfinite_cexp) blanchet@48975: blanchet@48975: lemma card_of_Sigma_ordLeq_Cinfinite: blanchet@48975: "\Cinfinite r; |I| \o r; \i \ I. |A i| \o r\ \ |SIGMA i : I. A i| \o r" blanchet@48975: unfolding cinfinite_def by (blast intro: card_of_Sigma_ordLeq_infinite_Field) blanchet@48975: traytel@54794: lemma card_order_cexp: traytel@54794: assumes "card_order r1" "card_order r2" traytel@54794: shows "card_order (r1 ^c r2)" traytel@54794: proof - traytel@54794: have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto traytel@54980: thus ?thesis unfolding cexp_def Func_def by simp traytel@54794: qed traytel@54794: traytel@54794: lemma Cinfinite_ordLess_cexp: traytel@54794: assumes r: "Cinfinite r" traytel@54794: shows "r o r ^c r" traytel@54794: by (rule cexp_mono1[OF ctwo_ordLeq_Cinfinite]) (auto simp: r ctwo_not_czero Card_order_ctwo) traytel@54794: finally show ?thesis . traytel@54794: qed traytel@54794: traytel@54794: lemma infinite_ordLeq_cexp: traytel@54794: assumes "Cinfinite r" traytel@54794: shows "r \o r ^c r" traytel@54794: by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]]) traytel@54794: traytel@54980: lemma czero_cexp: "Cnotzero r \ czero ^c r =o czero" traytel@54980: by (drule Cnotzero_imp_not_empty) (simp add: cexp_def czero_def card_of_empty_ordIso) traytel@54980: blanchet@58127: lemma Func_singleton: traytel@54980: fixes x :: 'b and A :: "'a set" traytel@54980: shows "|Func A {x}| =o |{x}|" traytel@54980: proof (rule ordIso_symmetric) traytel@54980: def f \ "\y a. if y = x \ a \ A then x else undefined" traytel@54980: have "Func A {x} \ f ` {x}" unfolding f_def Func_def by (force simp: fun_eq_iff) traytel@54980: hence "bij_betw f {x} (Func A {x})" unfolding bij_betw_def inj_on_def f_def Func_def traytel@54980: by (auto split: split_if_asm) traytel@54980: thus "|{x}| =o |Func A {x}|" using card_of_ordIso by blast traytel@54980: qed traytel@54980: traytel@54980: lemma cone_cexp: "cone ^c r =o cone" traytel@54980: unfolding cexp_def cone_def Field_card_of by (rule Func_singleton) traytel@54980: traytel@54980: lemma card_of_Func_squared: traytel@54980: fixes A :: "'a set" traytel@54980: shows "|Func (UNIV :: bool set) A| =o |A \ A|" traytel@54980: proof (rule ordIso_symmetric) traytel@54980: def f \ "\(x::'a,y) b. if A = {} then undefined else if b then x else y" traytel@54980: have "Func (UNIV :: bool set) A \ f ` (A \ A)" unfolding f_def Func_def traytel@54980: by (auto simp: image_iff fun_eq_iff split: option.splits split_if_asm) blast traytel@54980: hence "bij_betw f (A \ A) (Func (UNIV :: bool set) A)" traytel@54980: unfolding bij_betw_def inj_on_def f_def Func_def by (auto simp: fun_eq_iff) traytel@54980: thus "|A \ A| =o |Func (UNIV :: bool set) A|" using card_of_ordIso by blast traytel@54980: qed traytel@54980: traytel@54980: lemma cexp_ctwo: "r ^c ctwo =o r *c r" traytel@54980: unfolding cexp_def ctwo_def cprod_def Field_card_of by (rule card_of_Func_squared) traytel@54980: traytel@54980: lemma card_of_Func_Plus: traytel@54980: fixes A :: "'a set" and B :: "'b set" and C :: "'c set" traytel@54980: shows "|Func (A <+> B) C| =o |Func A C \ Func B C|" traytel@54980: proof (rule ordIso_symmetric) traytel@54980: def f \ "\(g :: 'a => 'c, h::'b \ 'c) ab. case ab of Inl a \ g a | Inr b \ h b" traytel@54980: def f' \ "\(f :: ('a + 'b) \ 'c). (\a. f (Inl a), \b. f (Inr b))" traytel@54980: have "f ` (Func A C \ Func B C) \ Func (A <+> B) C" traytel@54980: unfolding Func_def f_def by (force split: sum.splits) traytel@54980: moreover have "f' ` Func (A <+> B) C \ Func A C \ Func B C" unfolding Func_def f'_def by force traytel@54980: moreover have "\a \ Func A C \ Func B C. f' (f a) = a" unfolding f'_def f_def Func_def by auto traytel@54980: moreover have "\a' \ Func (A <+> B) C. f (f' a') = a'" unfolding f'_def f_def Func_def traytel@54980: by (auto split: sum.splits) traytel@54980: ultimately have "bij_betw f (Func A C \ Func B C) (Func (A <+> B) C)" traytel@54980: by (intro bij_betw_byWitness[of _ f' f]) traytel@54980: thus "|Func A C \ Func B C| =o |Func (A <+> B) C|" using card_of_ordIso by blast traytel@54980: qed traytel@54980: traytel@54980: lemma cexp_csum: "r ^c (s +c t) =o r ^c s *c r ^c t" traytel@54980: unfolding cexp_def cprod_def csum_def Field_card_of by (rule card_of_Func_Plus) traytel@54980: traytel@54794: blanchet@48975: subsection {* Powerset *} blanchet@48975: traytel@54794: definition cpow where "cpow r = |Pow (Field r)|" traytel@54794: traytel@54794: lemma card_order_cpow: "card_order r \ card_order (cpow r)" traytel@54794: by (simp only: cpow_def Field_card_order Pow_UNIV card_of_card_order_on) traytel@54794: traytel@54794: lemma cpow_greater_eq: "Card_order r \ r \o cpow r" traytel@54794: by (rule ordLess_imp_ordLeq) (simp only: cpow_def Card_order_Pow) traytel@54794: traytel@54794: lemma Cinfinite_cpow: "Cinfinite r \ Cinfinite (cpow r)" traytel@54794: unfolding cpow_def cinfinite_def by (metis Field_card_of card_of_Card_order infinite_Pow) traytel@54794: blanchet@48975: lemma Card_order_cpow: "Card_order (cpow r)" blanchet@48975: unfolding cpow_def by (rule card_of_Card_order) blanchet@48975: blanchet@48975: lemma cardSuc_ordLeq_cpow: "Card_order r \ cardSuc r \o cpow r" blanchet@48975: unfolding cpow_def by (metis Card_order_Pow cardSuc_ordLess_ordLeq card_of_Card_order) blanchet@48975: blanchet@48975: lemma cpow_cexp_ctwo: "cpow r =o ctwo ^c r" blanchet@48975: unfolding cpow_def ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func) blanchet@48975: traytel@54980: subsection {* Inverse image *} traytel@54980: traytel@54980: lemma vimage_ordLeq: traytel@54980: assumes "|A| \o k" and "\ a \ A. |vimage f {a}| \o k" and "Cinfinite k" traytel@54980: shows "|vimage f A| \o k" traytel@54980: proof- wenzelm@60585: have "vimage f A = (\a \ A. vimage f {a})" by auto wenzelm@60585: also have "|\a \ A. vimage f {a}| \o k" traytel@54980: using UNION_Cinfinite_bound[OF assms] . traytel@54980: finally show ?thesis . traytel@54980: qed traytel@54980: traytel@54980: subsection {* Maximum *} traytel@54980: traytel@54980: definition cmax where traytel@54980: "cmax r s = traytel@54980: (if cinfinite r \ cinfinite s then czero +c r +c s traytel@54980: else natLeq_on (max (card (Field r)) (card (Field s))) +c czero)" traytel@54980: traytel@54980: lemma cmax_com: "cmax r s =o cmax s r" traytel@54980: unfolding cmax_def traytel@54980: by (auto simp: max.commute intro: csum_cong2[OF csum_com] csum_cong2[OF czero_ordIso]) traytel@54980: traytel@54980: lemma cmax1: traytel@54980: assumes "Card_order r" "Card_order s" "s \o r" traytel@54980: shows "cmax r s =o r" traytel@54980: unfolding cmax_def proof (split if_splits, intro conjI impI) traytel@54980: assume "cinfinite r \ cinfinite s" traytel@54980: hence Cinf: "Cinfinite r" using assms(1,3) by (metis cinfinite_mono) traytel@54980: have "czero +c r +c s =o r +c s" by (rule csum_czero2[OF Card_order_csum]) traytel@54980: also have "r +c s =o r" by (rule csum_absorb1[OF Cinf assms(3)]) traytel@54980: finally show "czero +c r +c s =o r" . traytel@54980: next traytel@54980: assume "\ (cinfinite r \ cinfinite s)" traytel@54980: hence fin: "finite (Field r)" and "finite (Field s)" unfolding cinfinite_def by simp_all traytel@54980: moreover traytel@54980: { from assms(2) have "|Field s| =o s" by (rule card_of_Field_ordIso) traytel@54980: also from assms(3) have "s \o r" . traytel@54980: also from assms(1) have "r =o |Field r|" by (rule ordIso_symmetric[OF card_of_Field_ordIso]) traytel@54980: finally have "|Field s| \o |Field r|" . traytel@54980: } traytel@54980: ultimately have "card (Field s) \ card (Field r)" by (subst sym[OF finite_card_of_iff_card2]) traytel@54980: hence "max (card (Field r)) (card (Field s)) = card (Field r)" by (rule max_absorb1) traytel@54980: hence "natLeq_on (max (card (Field r)) (card (Field s))) +c czero = traytel@54980: natLeq_on (card (Field r)) +c czero" by simp traytel@54980: also have "\ =o natLeq_on (card (Field r))" by (rule csum_czero1[OF natLeq_on_Card_order]) traytel@54980: also have "natLeq_on (card (Field r)) =o |Field r|" traytel@54980: by (rule ordIso_symmetric[OF finite_imp_card_of_natLeq_on[OF fin]]) traytel@54980: also from assms(1) have "|Field r| =o r" by (rule card_of_Field_ordIso) traytel@54980: finally show "natLeq_on (max (card (Field r)) (card (Field s))) +c czero =o r" . traytel@54980: qed traytel@54980: traytel@54980: lemma cmax2: traytel@54980: assumes "Card_order r" "Card_order s" "r \o s" traytel@54980: shows "cmax r s =o s" traytel@54980: by (metis assms cmax1 cmax_com ordIso_transitive) traytel@54980: traytel@54980: lemma csum_absorb2: "Cinfinite r2 \ r1 \o r2 \ r1 +c r2 =o r2" traytel@54980: by (metis csum_absorb2') traytel@54980: traytel@54980: lemma cprod_infinite2': "\Cnotzero r1; Cinfinite r2; r1 \o r2\ \ r1 *c r2 =o r2" traytel@54980: unfolding ordIso_iff_ordLeq traytel@54980: by (intro conjI cprod_cinfinite_bound ordLeq_cprod2 ordLeq_refl) traytel@54980: (auto dest!: ordIso_imp_ordLeq not_ordLeq_ordLess simp: czero_def Card_order_empty) traytel@54980: traytel@54980: context traytel@54980: fixes r s traytel@54980: assumes r: "Cinfinite r" traytel@54980: and s: "Cinfinite s" traytel@54980: begin traytel@54980: traytel@54980: lemma cmax_csum: "cmax r s =o r +c s" traytel@54980: proof (cases "r \o s") traytel@54980: case True traytel@54980: hence "cmax r s =o s" by (metis cmax2 r s) traytel@54980: also have "s =o r +c s" by (metis True csum_absorb2 ordIso_symmetric s) traytel@54980: finally show ?thesis . traytel@54980: next traytel@54980: case False traytel@54980: hence "s \o r" by (metis ordLeq_total r s card_order_on_def) traytel@54980: hence "cmax r s =o r" by (metis cmax1 r s) traytel@54980: also have "r =o r +c s" by (metis `s \o r` csum_absorb1 ordIso_symmetric r) traytel@54980: finally show ?thesis . traytel@54980: qed traytel@54980: traytel@54980: lemma cmax_cprod: "cmax r s =o r *c s" traytel@54980: proof (cases "r \o s") traytel@54980: case True traytel@54980: hence "cmax r s =o s" by (metis cmax2 r s) traytel@54980: also have "s =o r *c s" by (metis Cinfinite_Cnotzero True cprod_infinite2' ordIso_symmetric r s) traytel@54980: finally show ?thesis . traytel@54980: next traytel@54980: case False traytel@54980: hence "s \o r" by (metis ordLeq_total r s card_order_on_def) traytel@54980: hence "cmax r s =o r" by (metis cmax1 r s) traytel@54980: also have "r =o r *c s" by (metis Cinfinite_Cnotzero `s \o r` cprod_infinite1' ordIso_symmetric r s) traytel@54980: finally show ?thesis . traytel@54980: qed traytel@54980: blanchet@48975: end traytel@54980: traytel@55174: lemma Card_order_cmax: traytel@55174: assumes r: "Card_order r" and s: "Card_order s" traytel@55174: shows "Card_order (cmax r s)" traytel@55174: unfolding cmax_def by (auto simp: Card_order_csum) traytel@55174: traytel@55174: lemma ordLeq_cmax: traytel@55174: assumes r: "Card_order r" and s: "Card_order s" traytel@55174: shows "r \o cmax r s \ s \o cmax r s" traytel@55174: proof- traytel@55174: {assume "r \o s" traytel@55174: hence ?thesis by (metis cmax2 ordIso_iff_ordLeq ordLeq_transitive r s) traytel@55174: } traytel@55174: moreover traytel@55174: {assume "s \o r" traytel@55174: hence ?thesis using cmax_com by (metis cmax2 ordIso_iff_ordLeq ordLeq_transitive r s) traytel@55174: } traytel@55174: ultimately show ?thesis using r s ordLeq_total unfolding card_order_on_def by auto traytel@55174: qed traytel@55174: traytel@55174: lemmas ordLeq_cmax1 = ordLeq_cmax[THEN conjunct1] and traytel@55174: ordLeq_cmax2 = ordLeq_cmax[THEN conjunct2] traytel@55174: traytel@55174: lemma finite_cmax: traytel@55174: assumes r: "Card_order r" and s: "Card_order s" traytel@55174: shows "finite (Field (cmax r s)) \ finite (Field r) \ finite (Field s)" traytel@55174: proof- traytel@55174: {assume "r \o s" traytel@55174: hence ?thesis by (metis cmax2 ordIso_finite_Field ordLeq_finite_Field r s) traytel@55174: } traytel@55174: moreover traytel@55174: {assume "s \o r" traytel@55174: hence ?thesis by (metis cmax1 ordIso_finite_Field ordLeq_finite_Field r s) traytel@55174: } traytel@55174: ultimately show ?thesis using r s ordLeq_total unfolding card_order_on_def by auto traytel@55174: qed traytel@55174: traytel@54980: end