# HG changeset patch # User blanchet # Date 1384794285 -3600 # Node ID 5c9819d7713be2260ab2cfa538520786d70c7645 # Parent 57e781b711b55d6438a35f2387c01ff422356cef compile diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/BNF/BNF_Util.thy --- a/src/HOL/BNF/BNF_Util.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/BNF/BNF_Util.thy Mon Nov 18 18:04:45 2013 +0100 @@ -9,7 +9,7 @@ header {* Library for Bounded Natural Functors *} theory BNF_Util -imports "../Cardinals/Cardinal_Arithmetic" +imports "../Cardinals/Cardinal_Arithmetic_FP" begin lemma subset_Collect_iff: "B \ A \ (B \ {x \ A. P x}) = (\x \ B. P x)" diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Cardinal_Arithmetic.thy --- a/src/HOL/Cardinals/Cardinal_Arithmetic.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* Cardinal Arithmetic *} theory Cardinal_Arithmetic -imports Cardinal_Arithmetic_LFP Cardinal_Order_Relation +imports Cardinal_Arithmetic_FP Cardinal_Order_Relation begin @@ -177,9 +177,6 @@ lemma Card_order_cpow: "Card_order (cpow r)" unfolding cpow_def by (rule card_of_Card_order) -lemma Cinfinite_cpow: "Cinfinite r \ Cinfinite (cpow r)" -unfolding cpow_def cinfinite_def by (metis Field_card_of card_of_Card_order infinite_Pow) - lemma cardSuc_ordLeq_cpow: "Card_order r \ cardSuc r \o cpow r" unfolding cpow_def by (metis Card_order_Pow cardSuc_ordLess_ordLeq card_of_Card_order) diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Cardinal_Arithmetic_FP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Cardinals/Cardinal_Arithmetic_FP.thy Mon Nov 18 18:04:45 2013 +0100 @@ -0,0 +1,743 @@ +(* Title: HOL/Cardinals/Cardinal_Arithmetic_FP.thy + Author: Dmitriy Traytel, TU Muenchen + Copyright 2012 + +Cardinal arithmetic (FP). +*) + +header {* Cardinal Arithmetic (FP) *} + +theory Cardinal_Arithmetic_FP +imports Cardinal_Order_Relation_FP +begin + +(*library candidate*) +lemma dir_image: "\\x y. (f x = f y) = (x = y); Card_order r\ \ r =o dir_image r f" +by (rule dir_image_ordIso) (auto simp add: inj_on_def card_order_on_def) + +(*should supersede a weaker lemma from the library*) +lemma dir_image_Field: "Field (dir_image r f) = f ` Field r" +unfolding dir_image_def Field_def Range_def Domain_def by fastforce + +lemma card_order_dir_image: + assumes bij: "bij f" and co: "card_order r" + shows "card_order (dir_image r f)" +proof - + from assms have "Field (dir_image r f) = UNIV" + using card_order_on_Card_order[of UNIV r] unfolding bij_def dir_image_Field by auto + moreover from bij have "\x y. (f x = f y) = (x = y)" unfolding bij_def inj_on_def by auto + with co have "Card_order (dir_image r f)" + using card_order_on_Card_order[of UNIV r] Card_order_ordIso2[OF _ dir_image] by blast + ultimately show ?thesis by auto +qed + +(*library candidate*) +lemma ordIso_refl: "Card_order r \ r =o r" +by (rule card_order_on_ordIso) + +(*library candidate*) +lemma ordLeq_refl: "Card_order r \ r \o r" +by (rule ordIso_imp_ordLeq, rule card_order_on_ordIso) + +(*library candidate*) +lemma card_of_ordIso_subst: "A = B \ |A| =o |B|" +by (simp only: ordIso_refl card_of_Card_order) + +(*library candidate*) +lemma Field_card_order: "card_order r \ Field r = UNIV" +using card_order_on_Card_order[of UNIV r] by simp + +(*library candidate*) +lemma card_of_Times_Plus_distrib: + "|A <*> (B <+> C)| =o |A <*> B <+> A <*> C|" (is "|?RHS| =o |?LHS|") +proof - + let ?f = "\(a, bc). case bc of Inl b \ Inl (a, b) | Inr c \ Inr (a, c)" + have "bij_betw ?f ?RHS ?LHS" unfolding bij_betw_def inj_on_def by force + thus ?thesis using card_of_ordIso by blast +qed + +(*library candidate*) +lemma Func_Times_Range: + "|Func A (B <*> C)| =o |Func A B <*> Func A C|" (is "|?LHS| =o |?RHS|") +proof - + let ?F = "\fg. (\x. if x \ A then fst (fg x) else undefined, + \x. if x \ A then snd (fg x) else undefined)" + let ?G = "\(f, g) x. if x \ A then (f x, g x) else undefined" + have "bij_betw ?F ?LHS ?RHS" unfolding bij_betw_def inj_on_def + proof safe + fix f g assume "f \ Func A B" "g \ Func A C" + thus "(f, g) \ ?F ` Func A (B \ C)" + by (intro image_eqI[of _ _ "?G (f, g)"]) (auto simp: Func_def) + qed (auto simp: Func_def fun_eq_iff, metis pair_collapse) + thus ?thesis using card_of_ordIso by blast +qed + + +subsection {* Zero *} + +definition czero where + "czero = card_of {}" + +lemma czero_ordIso: + "czero =o czero" +using card_of_empty_ordIso by (simp add: czero_def) + +lemma card_of_ordIso_czero_iff_empty: + "|A| =o (czero :: 'b rel) \ A = ({} :: 'a set)" +unfolding czero_def by (rule iffI[OF card_of_empty2]) (auto simp: card_of_refl card_of_empty_ordIso) + +(* A "not czero" Cardinal predicate *) +abbreviation Cnotzero where + "Cnotzero (r :: 'a rel) \ \(r =o (czero :: 'a rel)) \ Card_order r" + +(*helper*) +lemma Cnotzero_imp_not_empty: "Cnotzero r \ Field r \ {}" +by (metis Card_order_iff_ordIso_card_of czero_def) + +lemma czeroI: + "\Card_order r; Field r = {}\ \ r =o czero" +using Cnotzero_imp_not_empty ordIso_transitive[OF _ czero_ordIso] by blast + +lemma czeroE: + "r =o czero \ Field r = {}" +unfolding czero_def +by (drule card_of_cong) (simp only: Field_card_of card_of_empty2) + +lemma Cnotzero_mono: + "\Cnotzero r; Card_order q; r \o q\ \ Cnotzero q" +apply (rule ccontr) +apply auto +apply (drule czeroE) +apply (erule notE) +apply (erule czeroI) +apply (drule card_of_mono2) +apply (simp only: card_of_empty3) +done + +subsection {* (In)finite cardinals *} + +definition cinfinite where + "cinfinite r = infinite (Field r)" + +abbreviation Cinfinite where + "Cinfinite r \ cinfinite r \ Card_order r" + +definition cfinite where + "cfinite r = finite (Field r)" + +abbreviation Cfinite where + "Cfinite r \ cfinite r \ Card_order r" + +lemma Cfinite_ordLess_Cinfinite: "\Cfinite r; Cinfinite s\ \ r o r" +proof - + from inf have "natLeq \o |Field r|" by (simp add: cinfinite_def infinite_iff_natLeq_ordLeq) + also from inf have "|Field r| =o r" by (simp add: card_of_unique ordIso_symmetric) + finally show ?thesis . +qed + +lemma cinfinite_not_czero: "cinfinite r \ \ (r =o (czero :: 'a rel))" +unfolding cinfinite_def by (metis czeroE finite.emptyI) + +lemma Cinfinite_Cnotzero: "Cinfinite r \ Cnotzero r" +by (metis cinfinite_not_czero) + +lemma Cinfinite_cong: "\r1 =o r2; Cinfinite r1\ \ Cinfinite r2" +by (metis Card_order_ordIso2 card_of_mono2 card_of_ordLeq_infinite cinfinite_def ordIso_iff_ordLeq) + +lemma cinfinite_mono: "\r1 \o r2; cinfinite r1\ \ cinfinite r2" +by (metis card_of_mono2 card_of_ordLeq_infinite cinfinite_def) + + +subsection {* Binary sum *} + +definition csum (infixr "+c" 65) where + "r1 +c r2 \ |Field r1 <+> Field r2|" + +lemma Field_csum: "Field (r +c s) = Inl ` Field r \ Inr ` Field s" + unfolding csum_def Field_card_of by auto + +lemma Card_order_csum: + "Card_order (r1 +c r2)" +unfolding csum_def by (simp add: card_of_Card_order) + +lemma csum_Cnotzero1: + "Cnotzero r1 \ Cnotzero (r1 +c r2)" +unfolding csum_def +by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE) + +lemma card_order_csum: + assumes "card_order r1" "card_order r2" + shows "card_order (r1 +c r2)" +proof - + have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto + thus ?thesis unfolding csum_def by (auto simp: card_of_card_order_on) +qed + +lemma cinfinite_csum: + "cinfinite r1 \ cinfinite r2 \ cinfinite (r1 +c r2)" +unfolding cinfinite_def csum_def by (auto simp: Field_card_of) + +lemma Cinfinite_csum1: + "Cinfinite r1 \ Cinfinite (r1 +c r2)" +unfolding cinfinite_def csum_def by (metis Field_card_of card_of_Card_order finite_Plus_iff) + +lemma Cinfinite_csum: + "Cinfinite r1 \ Cinfinite r2 \ Cinfinite (r1 +c r2)" +unfolding cinfinite_def csum_def by (metis Field_card_of card_of_Card_order finite_Plus_iff) + +lemma Cinfinite_csum_strong: + "\Cinfinite r1; Cinfinite r2\ \ Cinfinite (r1 +c r2)" +by (metis Cinfinite_csum) + +lemma csum_cong: "\p1 =o r1; p2 =o r2\ \ p1 +c p2 =o r1 +c r2" +by (simp only: csum_def ordIso_Plus_cong) + +lemma csum_cong1: "p1 =o r1 \ p1 +c q =o r1 +c q" +by (simp only: csum_def ordIso_Plus_cong1) + +lemma csum_cong2: "p2 =o r2 \ q +c p2 =o q +c r2" +by (simp only: csum_def ordIso_Plus_cong2) + +lemma csum_mono: "\p1 \o r1; p2 \o r2\ \ p1 +c p2 \o r1 +c r2" +by (simp only: csum_def ordLeq_Plus_mono) + +lemma csum_mono1: "p1 \o r1 \ p1 +c q \o r1 +c q" +by (simp only: csum_def ordLeq_Plus_mono1) + +lemma csum_mono2: "p2 \o r2 \ q +c p2 \o q +c r2" +by (simp only: csum_def ordLeq_Plus_mono2) + +lemma ordLeq_csum1: "Card_order p1 \ p1 \o p1 +c p2" +by (simp only: csum_def Card_order_Plus1) + +lemma ordLeq_csum2: "Card_order p2 \ p2 \o p1 +c p2" +by (simp only: csum_def Card_order_Plus2) + +lemma csum_com: "p1 +c p2 =o p2 +c p1" +by (simp only: csum_def card_of_Plus_commute) + +lemma csum_assoc: "(p1 +c p2) +c p3 =o p1 +c p2 +c p3" +by (simp only: csum_def Field_card_of card_of_Plus_assoc) + +lemma Cfinite_csum: "\Cfinite r; Cfinite s\ \ Cfinite (r +c s)" + unfolding cfinite_def csum_def Field_card_of using card_of_card_order_on by simp + +lemma csum_csum: "(r1 +c r2) +c (r3 +c r4) =o (r1 +c r3) +c (r2 +c r4)" +proof - + have "(r1 +c r2) +c (r3 +c r4) =o r1 +c r2 +c (r3 +c r4)" + by (metis csum_assoc) + also have "r1 +c r2 +c (r3 +c r4) =o r1 +c (r2 +c r3) +c r4" + by (metis csum_assoc csum_cong2 ordIso_symmetric) + also have "r1 +c (r2 +c r3) +c r4 =o r1 +c (r3 +c r2) +c r4" + by (metis csum_com csum_cong1 csum_cong2) + also have "r1 +c (r3 +c r2) +c r4 =o r1 +c r3 +c r2 +c r4" + by (metis csum_assoc csum_cong2 ordIso_symmetric) + also have "r1 +c r3 +c r2 +c r4 =o (r1 +c r3) +c (r2 +c r4)" + by (metis csum_assoc ordIso_symmetric) + finally show ?thesis . +qed + +lemma Plus_csum: "|A <+> B| =o |A| +c |B|" +by (simp only: csum_def Field_card_of card_of_refl) + +lemma Un_csum: "|A \ B| \o |A| +c |B|" +using ordLeq_ordIso_trans[OF card_of_Un_Plus_ordLeq Plus_csum] by blast + + +subsection {* One *} + +definition cone where + "cone = card_of {()}" + +lemma Card_order_cone: "Card_order cone" +unfolding cone_def by (rule card_of_Card_order) + +lemma Cfinite_cone: "Cfinite cone" + unfolding cfinite_def by (simp add: Card_order_cone) + +lemma cone_not_czero: "\ (cone =o czero)" +unfolding czero_def cone_def by (metis empty_not_insert card_of_empty3[of "{()}"] ordIso_iff_ordLeq) + +lemma cone_ordLeq_Cnotzero: "Cnotzero r \ cone \o r" +unfolding cone_def by (metis Card_order_singl_ordLeq czeroI) + + +subsection{* Two *} + +definition ctwo where + "ctwo = |UNIV :: bool set|" + +lemma Card_order_ctwo: "Card_order ctwo" +unfolding ctwo_def by (rule card_of_Card_order) + +lemma ctwo_not_czero: "\ (ctwo =o czero)" +using card_of_empty3[of "UNIV :: bool set"] ordIso_iff_ordLeq +unfolding czero_def ctwo_def by (metis UNIV_not_empty) + +lemma ctwo_Cnotzero: "Cnotzero ctwo" +by (simp add: ctwo_not_czero Card_order_ctwo) + + +subsection {* Family sum *} + +definition Csum where + "Csum r rs \ |SIGMA i : Field r. Field (rs i)|" + +(* Similar setup to the one for SIGMA from theory Big_Operators: *) +syntax "_Csum" :: + "pttrn => ('a * 'a) set => 'b * 'b set => (('a * 'b) * ('a * 'b)) set" + ("(3CSUM _:_. _)" [0, 51, 10] 10) + +translations + "CSUM i:r. rs" == "CONST Csum r (%i. rs)" + +lemma SIGMA_CSUM: "|SIGMA i : I. As i| = (CSUM i : |I|. |As i| )" +by (auto simp: Csum_def Field_card_of) + +(* NB: Always, under the cardinal operator, +operations on sets are reduced automatically to operations on cardinals. +This should make cardinal reasoning more direct and natural. *) + + +subsection {* Product *} + +definition cprod (infixr "*c" 80) where + "r1 *c r2 = |Field r1 <*> Field r2|" + +lemma card_order_cprod: + assumes "card_order r1" "card_order r2" + shows "card_order (r1 *c r2)" +proof - + have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto + thus ?thesis by (auto simp: cprod_def card_of_card_order_on) +qed + +lemma Card_order_cprod: "Card_order (r1 *c r2)" +by (simp only: cprod_def Field_card_of card_of_card_order_on) + +lemma cprod_mono1: "p1 \o r1 \ p1 *c q \o r1 *c q" +by (simp only: cprod_def ordLeq_Times_mono1) + +lemma cprod_mono2: "p2 \o r2 \ q *c p2 \o q *c r2" +by (simp only: cprod_def ordLeq_Times_mono2) + +lemma ordLeq_cprod2: "\Cnotzero p1; Card_order p2\ \ p2 \o p1 *c p2" +unfolding cprod_def by (metis Card_order_Times2 czeroI) + +lemma cinfinite_cprod: "\cinfinite r1; cinfinite r2\ \ cinfinite (r1 *c r2)" +by (simp add: cinfinite_def cprod_def Field_card_of infinite_cartesian_product) + +lemma cinfinite_cprod2: "\Cnotzero r1; Cinfinite r2\ \ cinfinite (r1 *c r2)" +by (metis cinfinite_mono ordLeq_cprod2) + +lemma Cinfinite_cprod2: "\Cnotzero r1; Cinfinite r2\ \ Cinfinite (r1 *c r2)" +by (blast intro: cinfinite_cprod2 Card_order_cprod) + +lemma cprod_com: "p1 *c p2 =o p2 *c p1" +by (simp only: cprod_def card_of_Times_commute) + +lemma card_of_Csum_Times: + "\i \ I. |A i| \o |B| \ (CSUM i : |I|. |A i| ) \o |I| *c |B|" +by (simp only: Csum_def cprod_def Field_card_of card_of_Sigma_Times) + +lemma card_of_Csum_Times': + assumes "Card_order r" "\i \ I. |A i| \o r" + shows "(CSUM i : |I|. |A i| ) \o |I| *c r" +proof - + from assms(1) have *: "r =o |Field r|" by (simp add: card_of_unique) + with assms(2) have "\i \ I. |A i| \o |Field r|" by (blast intro: ordLeq_ordIso_trans) + hence "(CSUM i : |I|. |A i| ) \o |I| *c |Field r|" by (simp only: card_of_Csum_Times) + also from * have "|I| *c |Field r| \o |I| *c r" + by (simp only: Field_card_of card_of_refl cprod_def ordIso_imp_ordLeq) + finally show ?thesis . +qed + +lemma cprod_csum_distrib1: "r1 *c r2 +c r1 *c r3 =o r1 *c (r2 +c r3)" +unfolding csum_def cprod_def by (simp add: Field_card_of card_of_Times_Plus_distrib ordIso_symmetric) + +lemma csum_absorb2': "\Card_order r2; r1 \o r2; cinfinite r1 \ cinfinite r2\ \ r1 +c r2 =o r2" +unfolding csum_def by (metis Card_order_Plus_infinite cinfinite_def cinfinite_mono) + +lemma csum_absorb1': + assumes card: "Card_order r2" + and r12: "r1 \o r2" and cr12: "cinfinite r1 \ cinfinite r2" + shows "r2 +c r1 =o r2" +by (rule ordIso_transitive, rule csum_com, rule csum_absorb2', (simp only: assms)+) + +lemma csum_absorb1: "\Cinfinite r2; r1 \o r2\ \ r2 +c r1 =o r2" +by (rule csum_absorb1') auto + + +subsection {* Exponentiation *} + +definition cexp (infixr "^c" 90) where + "r1 ^c r2 \ |Func (Field r2) (Field r1)|" + +lemma Card_order_cexp: "Card_order (r1 ^c r2)" +unfolding cexp_def by (rule card_of_Card_order) + +lemma cexp_mono': + assumes 1: "p1 \o r1" and 2: "p2 \o r2" + and n: "Field p2 = {} \ Field r2 = {}" + shows "p1 ^c p2 \o r1 ^c r2" +proof(cases "Field p1 = {}") + case True + hence "|Field |Func (Field p2) (Field p1)|| \o cone" + unfolding cone_def Field_card_of + by (cases "Field p2 = {}", auto intro: card_of_ordLeqI2 simp: Func_empty) + (metis Func_is_emp card_of_empty ex_in_conv) + hence "|Func (Field p2) (Field p1)| \o cone" by (simp add: Field_card_of cexp_def) + hence "p1 ^c p2 \o cone" unfolding cexp_def . + thus ?thesis + proof (cases "Field p2 = {}") + case True + with n have "Field r2 = {}" . + hence "cone \o r1 ^c r2" unfolding cone_def cexp_def Func_def by (auto intro: card_of_ordLeqI) + thus ?thesis using `p1 ^c p2 \o cone` ordLeq_transitive by auto + next + case False with True have "|Field (p1 ^c p2)| =o czero" + unfolding card_of_ordIso_czero_iff_empty cexp_def Field_card_of Func_def by auto + thus ?thesis unfolding cexp_def card_of_ordIso_czero_iff_empty Field_card_of + by (simp add: card_of_empty) + qed +next + case False + have 1: "|Field p1| \o |Field r1|" and 2: "|Field p2| \o |Field r2|" + using 1 2 by (auto simp: card_of_mono2) + obtain f1 where f1: "f1 ` Field r1 = Field p1" + using 1 unfolding card_of_ordLeq2[OF False, symmetric] by auto + obtain f2 where f2: "inj_on f2 (Field p2)" "f2 ` Field p2 \ Field r2" + using 2 unfolding card_of_ordLeq[symmetric] by blast + have 0: "Func_map (Field p2) f1 f2 ` (Field (r1 ^c r2)) = Field (p1 ^c p2)" + unfolding cexp_def Field_card_of using Func_map_surj[OF f1 f2 n, symmetric] . + have 00: "Field (p1 ^c p2) \ {}" unfolding cexp_def Field_card_of Func_is_emp + using False by simp + show ?thesis + using 0 card_of_ordLeq2[OF 00] unfolding cexp_def Field_card_of by blast +qed + +lemma cexp_mono: + assumes 1: "p1 \o r1" and 2: "p2 \o r2" + and n: "p2 =o czero \ r2 =o czero" and card: "Card_order p2" + shows "p1 ^c p2 \o r1 ^c r2" + by (metis (full_types) "1" "2" card cexp_mono' czeroE czeroI n) + +lemma cexp_mono1: + assumes 1: "p1 \o r1" and q: "Card_order q" + shows "p1 ^c q \o r1 ^c q" +using ordLeq_refl[OF q] by (rule cexp_mono[OF 1]) (auto simp: q) + +lemma cexp_mono2': + assumes 2: "p2 \o r2" and q: "Card_order q" + and n: "Field p2 = {} \ Field r2 = {}" + shows "q ^c p2 \o q ^c r2" +using ordLeq_refl[OF q] by (rule cexp_mono'[OF _ 2 n]) auto + +lemma cexp_mono2: + assumes 2: "p2 \o r2" and q: "Card_order q" + and n: "p2 =o czero \ r2 =o czero" and card: "Card_order p2" + shows "q ^c p2 \o q ^c r2" +using ordLeq_refl[OF q] by (rule cexp_mono[OF _ 2 n card]) auto + +lemma cexp_mono2_Cnotzero: + assumes "p2 \o r2" "Card_order q" "Cnotzero p2" + shows "q ^c p2 \o q ^c r2" +by (metis assms cexp_mono2' czeroI) + +lemma cexp_cong: + assumes 1: "p1 =o r1" and 2: "p2 =o r2" + and Cr: "Card_order r2" + and Cp: "Card_order p2" + shows "p1 ^c p2 =o r1 ^c r2" +proof - + obtain f where "bij_betw f (Field p2) (Field r2)" + using 2 card_of_ordIso[of "Field p2" "Field r2"] card_of_cong by auto + hence 0: "Field p2 = {} \ Field r2 = {}" unfolding bij_betw_def by auto + have r: "p2 =o czero \ r2 =o czero" + and p: "r2 =o czero \ p2 =o czero" + using 0 Cr Cp czeroE czeroI by auto + show ?thesis using 0 1 2 unfolding ordIso_iff_ordLeq + using r p cexp_mono[OF _ _ _ Cp] cexp_mono[OF _ _ _ Cr] by blast +qed + +lemma cexp_cong1: + assumes 1: "p1 =o r1" and q: "Card_order q" + shows "p1 ^c q =o r1 ^c q" +by (rule cexp_cong[OF 1 _ q q]) (rule ordIso_refl[OF q]) + +lemma cexp_cong2: + assumes 2: "p2 =o r2" and q: "Card_order q" and p: "Card_order p2" + shows "q ^c p2 =o q ^c r2" +by (rule cexp_cong[OF _ 2]) (auto simp only: ordIso_refl Card_order_ordIso2[OF p 2] q p) + +lemma cexp_cone: + assumes "Card_order r" + shows "r ^c cone =o r" +proof - + have "r ^c cone =o |Field r|" + unfolding cexp_def cone_def Field_card_of Func_empty + card_of_ordIso[symmetric] bij_betw_def Func_def inj_on_def image_def + by (rule exI[of _ "\f. f ()"]) auto + also have "|Field r| =o r" by (rule card_of_Field_ordIso[OF assms]) + finally show ?thesis . +qed + +lemma cexp_cprod: + assumes r1: "Card_order r1" + shows "(r1 ^c r2) ^c r3 =o r1 ^c (r2 *c r3)" (is "?L =o ?R") +proof - + have "?L =o r1 ^c (r3 *c r2)" + unfolding cprod_def cexp_def Field_card_of + using card_of_Func_Times by(rule ordIso_symmetric) + also have "r1 ^c (r3 *c r2) =o ?R" + apply(rule cexp_cong2) using cprod_com r1 by (auto simp: Card_order_cprod) + finally show ?thesis . +qed + +lemma cprod_infinite1': "\Cinfinite r; Cnotzero p; p \o r\ \ r *c p =o r" +unfolding cinfinite_def cprod_def +by (rule Card_order_Times_infinite[THEN conjunct1]) (blast intro: czeroI)+ + +lemma cexp_cprod_ordLeq: + assumes r1: "Card_order r1" and r2: "Cinfinite r2" + and r3: "Cnotzero r3" "r3 \o r2" + shows "(r1 ^c r2) ^c r3 =o r1 ^c r2" (is "?L =o ?R") +proof- + have "?L =o r1 ^c (r2 *c r3)" using cexp_cprod[OF r1] . + also have "r1 ^c (r2 *c r3) =o ?R" + apply(rule cexp_cong2) + apply(rule cprod_infinite1'[OF r2 r3]) using r1 r2 by (fastforce simp: Card_order_cprod)+ + finally show ?thesis . +qed + +lemma Cnotzero_UNIV: "Cnotzero |UNIV|" +by (auto simp: card_of_Card_order card_of_ordIso_czero_iff_empty) + +lemma ordLess_ctwo_cexp: + assumes "Card_order r" + shows "r o q ^c r" +proof (cases "q =o (czero :: 'a rel)") + case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans) +next + case False + thus ?thesis + apply - + apply (rule ordIso_ordLeq_trans) + apply (rule ordIso_symmetric) + apply (rule cexp_cone) + apply (rule assms(2)) + apply (rule cexp_mono2) + apply (rule cone_ordLeq_Cnotzero) + apply (rule assms(1)) + apply (rule assms(2)) + apply (rule notE) + apply (rule cone_not_czero) + apply assumption + apply (rule Card_order_cone) + done +qed + +lemma ordLeq_cexp2: + assumes "ctwo \o q" "Card_order r" + shows "r \o q ^c r" +proof (cases "r =o (czero :: 'a rel)") + case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans) +next + case False thus ?thesis + apply - + apply (rule ordLess_imp_ordLeq) + apply (rule ordLess_ordLeq_trans) + apply (rule ordLess_ctwo_cexp) + apply (rule assms(2)) + apply (rule cexp_mono1) + apply (rule assms(1)) + apply (rule assms(2)) + done +qed + +lemma cinfinite_cexp: "\ctwo \o q; Cinfinite r\ \ cinfinite (q ^c r)" +by (metis assms cinfinite_mono ordLeq_cexp2) + +lemma Cinfinite_cexp: + "\ctwo \o q; Cinfinite r\ \ Cinfinite (q ^c r)" +by (simp add: cinfinite_cexp Card_order_cexp) + +lemma ctwo_ordLess_natLeq: "ctwo ctwo o r" +by (rule ordLess_imp_ordLeq[OF ctwo_ordLess_Cinfinite[OF assms]]) + +lemma Un_Cinfinite_bound: "\|A| \o r; |B| \o r; Cinfinite r\ \ |A \ B| \o r" +by (auto simp add: cinfinite_def card_of_Un_ordLeq_infinite_Field) + +lemma UNION_Cinfinite_bound: "\|I| \o r; \i \ I. |A i| \o r; Cinfinite r\ \ |\i \ I. A i| \o r" +by (auto simp add: card_of_UNION_ordLeq_infinite_Field cinfinite_def) + +lemma csum_cinfinite_bound: + assumes "p \o r" "q \o r" "Card_order p" "Card_order q" "Cinfinite r" + shows "p +c q \o r" +proof - + from assms(1-4) have "|Field p| \o r" "|Field q| \o r" + unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+ + with assms show ?thesis unfolding cinfinite_def csum_def + by (blast intro: card_of_Plus_ordLeq_infinite_Field) +qed + +lemma cprod_cinfinite_bound: + assumes "p \o r" "q \o r" "Card_order p" "Card_order q" "Cinfinite r" + shows "p *c q \o r" +proof - + from assms(1-4) have "|Field p| \o r" "|Field q| \o r" + unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+ + with assms show ?thesis unfolding cinfinite_def cprod_def + by (blast intro: card_of_Times_ordLeq_infinite_Field) +qed + +lemma cprod_csum_cexp: + "r1 *c r2 \o (r1 +c r2) ^c ctwo" +unfolding cprod_def csum_def cexp_def ctwo_def Field_card_of +proof - + let ?f = "\(a, b). %x. if x then Inl a else Inr b" + have "inj_on ?f (Field r1 \ Field r2)" (is "inj_on _ ?LHS") + by (auto simp: inj_on_def fun_eq_iff split: bool.split) + moreover + have "?f ` ?LHS \ Func (UNIV :: bool set) (Field r1 <+> Field r2)" (is "_ \ ?RHS") + by (auto simp: Func_def) + ultimately show "|?LHS| \o |?RHS|" using card_of_ordLeq by blast +qed + +lemma Cfinite_cprod_Cinfinite: "\Cfinite r; Cinfinite s\ \ r *c s \o s" +by (intro cprod_cinfinite_bound) + (auto intro: ordLeq_refl ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite]) + +lemma cprod_cexp: "(r *c s) ^c t =o r ^c t *c s ^c t" + unfolding cprod_def cexp_def Field_card_of by (rule Func_Times_Range) + +lemma cprod_cexp_csum_cexp_Cinfinite: + assumes t: "Cinfinite t" + shows "(r *c s) ^c t \o (r +c s) ^c t" +proof - + have "(r *c s) ^c t \o ((r +c s) ^c ctwo) ^c t" + by (rule cexp_mono1[OF cprod_csum_cexp conjunct2[OF t]]) + also have "((r +c s) ^c ctwo) ^c t =o (r +c s) ^c (ctwo *c t)" + by (rule cexp_cprod[OF Card_order_csum]) + also have "(r +c s) ^c (ctwo *c t) =o (r +c s) ^c (t *c ctwo)" + by (rule cexp_cong2[OF cprod_com Card_order_csum Card_order_cprod]) + also have "(r +c s) ^c (t *c ctwo) =o ((r +c s) ^c t) ^c ctwo" + by (rule ordIso_symmetric[OF cexp_cprod[OF Card_order_csum]]) + also have "((r +c s) ^c t) ^c ctwo =o (r +c s) ^c t" + by (rule cexp_cprod_ordLeq[OF Card_order_csum t ctwo_Cnotzero ctwo_ordLeq_Cinfinite[OF t]]) + finally show ?thesis . +qed + +lemma Cfinite_cexp_Cinfinite: + assumes s: "Cfinite s" and t: "Cinfinite t" + shows "s ^c t \o ctwo ^c t" +proof (cases "s \o ctwo") + case True thus ?thesis using t by (blast intro: cexp_mono1) +next + case False + hence "ctwo \o s" by (metis card_order_on_well_order_on ctwo_Cnotzero ordLeq_total s) + hence "Cnotzero s" by (metis Cnotzero_mono ctwo_Cnotzero s) + hence st: "Cnotzero (s *c t)" by (metis Cinfinite_cprod2 cinfinite_not_czero t) + have "s ^c t \o (ctwo ^c s) ^c t" + using assms by (blast intro: cexp_mono1 ordLess_imp_ordLeq[OF ordLess_ctwo_cexp]) + also have "(ctwo ^c s) ^c t =o ctwo ^c (s *c t)" + by (blast intro: Card_order_ctwo cexp_cprod) + also have "ctwo ^c (s *c t) \o ctwo ^c t" + using assms st by (intro cexp_mono2_Cnotzero Cfinite_cprod_Cinfinite Card_order_ctwo) + finally show ?thesis . +qed + +lemma csum_Cfinite_cexp_Cinfinite: + assumes r: "Card_order r" and s: "Cfinite s" and t: "Cinfinite t" + shows "(r +c s) ^c t \o (r +c ctwo) ^c t" +proof (cases "Cinfinite r") + case True + hence "r +c s =o r" by (intro csum_absorb1 ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite] s) + hence "(r +c s) ^c t =o r ^c t" using t by (blast intro: cexp_cong1) + also have "r ^c t \o (r +c ctwo) ^c t" using t by (blast intro: cexp_mono1 ordLeq_csum1 r) + finally show ?thesis . +next + case False + with r have "Cfinite r" unfolding cinfinite_def cfinite_def by auto + hence "Cfinite (r +c s)" by (intro Cfinite_csum s) + hence "(r +c s) ^c t \o ctwo ^c t" by (intro Cfinite_cexp_Cinfinite t) + also have "ctwo ^c t \o (r +c ctwo) ^c t" using t + by (blast intro: cexp_mono1 ordLeq_csum2 Card_order_ctwo) + finally show ?thesis . +qed + +lemma card_order_cexp: + assumes "card_order r1" "card_order r2" + shows "card_order (r1 ^c r2)" +proof - + have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto + thus ?thesis unfolding cexp_def Func_def by (simp add: card_of_card_order_on) +qed + +lemma Cinfinite_ordLess_cexp: + assumes r: "Cinfinite r" + shows "r o r ^c r" + by (rule cexp_mono1[OF ctwo_ordLeq_Cinfinite]) (auto simp: r ctwo_not_czero Card_order_ctwo) + finally show ?thesis . +qed + +lemma infinite_ordLeq_cexp: + assumes "Cinfinite r" + shows "r \o r ^c r" +by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]]) + +(* cardSuc *) + +lemma Cinfinite_cardSuc: "Cinfinite r \ Cinfinite (cardSuc r)" +by (simp add: cinfinite_def cardSuc_Card_order cardSuc_finite) + +lemma cardSuc_UNION_Cinfinite: + assumes "Cinfinite r" "relChain (cardSuc r) As" "B \ (UN i : Field (cardSuc r). As i)" "|B| <=o r" + shows "EX i : Field (cardSuc r). B \ As i" +using cardSuc_UNION assms unfolding cinfinite_def by blast + +subsection {* Powerset *} + +definition cpow where "cpow r = |Pow (Field r)|" + +lemma card_order_cpow: "card_order r \ card_order (cpow r)" +by (simp only: cpow_def Field_card_order Pow_UNIV card_of_card_order_on) + +lemma cpow_greater_eq: "Card_order r \ r \o cpow r" +by (rule ordLess_imp_ordLeq) (simp only: cpow_def Card_order_Pow) + +lemma Cinfinite_cpow: "Cinfinite r \ Cinfinite (cpow r)" +unfolding cpow_def cinfinite_def by (metis Field_card_of card_of_Card_order infinite_Pow) + +subsection {* Lists *} + +definition clists where "clists r = |lists (Field r)|" + +end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Cardinal_Arithmetic_LFP.thy --- a/src/HOL/Cardinals/Cardinal_Arithmetic_LFP.thy Mon Nov 18 18:04:45 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,740 +0,0 @@ -(* Title: HOL/Cardinals/Cardinal_Arithmetic_LFP.thy - Author: Dmitriy Traytel, TU Muenchen - Copyright 2012 - -Cardinal arithmetic. -*) - -header {* Cardinal Arithmetic *} - -theory Cardinal_Arithmetic_LFP -imports Cardinal_Order_Relation_LFP -begin - -(*library candidate*) -lemma dir_image: "\\x y. (f x = f y) = (x = y); Card_order r\ \ r =o dir_image r f" -by (rule dir_image_ordIso) (auto simp add: inj_on_def card_order_on_def) - -(*should supersede a weaker lemma from the library*) -lemma dir_image_Field: "Field (dir_image r f) = f ` Field r" -unfolding dir_image_def Field_def Range_def Domain_def by fastforce - -lemma card_order_dir_image: - assumes bij: "bij f" and co: "card_order r" - shows "card_order (dir_image r f)" -proof - - from assms have "Field (dir_image r f) = UNIV" - using card_order_on_Card_order[of UNIV r] unfolding bij_def dir_image_Field by auto - moreover from bij have "\x y. (f x = f y) = (x = y)" unfolding bij_def inj_on_def by auto - with co have "Card_order (dir_image r f)" - using card_order_on_Card_order[of UNIV r] Card_order_ordIso2[OF _ dir_image] by blast - ultimately show ?thesis by auto -qed - -(*library candidate*) -lemma ordIso_refl: "Card_order r \ r =o r" -by (rule card_order_on_ordIso) - -(*library candidate*) -lemma ordLeq_refl: "Card_order r \ r \o r" -by (rule ordIso_imp_ordLeq, rule card_order_on_ordIso) - -(*library candidate*) -lemma card_of_ordIso_subst: "A = B \ |A| =o |B|" -by (simp only: ordIso_refl card_of_Card_order) - -(*library candidate*) -lemma Field_card_order: "card_order r \ Field r = UNIV" -using card_order_on_Card_order[of UNIV r] by simp - -(*library candidate*) -lemma card_of_Times_Plus_distrib: - "|A <*> (B <+> C)| =o |A <*> B <+> A <*> C|" (is "|?RHS| =o |?LHS|") -proof - - let ?f = "\(a, bc). case bc of Inl b \ Inl (a, b) | Inr c \ Inr (a, c)" - have "bij_betw ?f ?RHS ?LHS" unfolding bij_betw_def inj_on_def by force - thus ?thesis using card_of_ordIso by blast -qed - -(*library candidate*) -lemma Func_Times_Range: - "|Func A (B <*> C)| =o |Func A B <*> Func A C|" (is "|?LHS| =o |?RHS|") -proof - - let ?F = "\fg. (\x. if x \ A then fst (fg x) else undefined, - \x. if x \ A then snd (fg x) else undefined)" - let ?G = "\(f, g) x. if x \ A then (f x, g x) else undefined" - have "bij_betw ?F ?LHS ?RHS" unfolding bij_betw_def inj_on_def - proof safe - fix f g assume "f \ Func A B" "g \ Func A C" - thus "(f, g) \ ?F ` Func A (B \ C)" - by (intro image_eqI[of _ _ "?G (f, g)"]) (auto simp: Func_def) - qed (auto simp: Func_def fun_eq_iff, metis pair_collapse) - thus ?thesis using card_of_ordIso by blast -qed - - -subsection {* Zero *} - -definition czero where - "czero = card_of {}" - -lemma czero_ordIso: - "czero =o czero" -using card_of_empty_ordIso by (simp add: czero_def) - -lemma card_of_ordIso_czero_iff_empty: - "|A| =o (czero :: 'b rel) \ A = ({} :: 'a set)" -unfolding czero_def by (rule iffI[OF card_of_empty2]) (auto simp: card_of_refl card_of_empty_ordIso) - -(* A "not czero" Cardinal predicate *) -abbreviation Cnotzero where - "Cnotzero (r :: 'a rel) \ \(r =o (czero :: 'a rel)) \ Card_order r" - -(*helper*) -lemma Cnotzero_imp_not_empty: "Cnotzero r \ Field r \ {}" -by (metis Card_order_iff_ordIso_card_of czero_def) - -lemma czeroI: - "\Card_order r; Field r = {}\ \ r =o czero" -using Cnotzero_imp_not_empty ordIso_transitive[OF _ czero_ordIso] by blast - -lemma czeroE: - "r =o czero \ Field r = {}" -unfolding czero_def -by (drule card_of_cong) (simp only: Field_card_of card_of_empty2) - -lemma Cnotzero_mono: - "\Cnotzero r; Card_order q; r \o q\ \ Cnotzero q" -apply (rule ccontr) -apply auto -apply (drule czeroE) -apply (erule notE) -apply (erule czeroI) -apply (drule card_of_mono2) -apply (simp only: card_of_empty3) -done - -subsection {* (In)finite cardinals *} - -definition cinfinite where - "cinfinite r = infinite (Field r)" - -abbreviation Cinfinite where - "Cinfinite r \ cinfinite r \ Card_order r" - -definition cfinite where - "cfinite r = finite (Field r)" - -abbreviation Cfinite where - "Cfinite r \ cfinite r \ Card_order r" - -lemma Cfinite_ordLess_Cinfinite: "\Cfinite r; Cinfinite s\ \ r o r" -proof - - from inf have "natLeq \o |Field r|" by (simp add: cinfinite_def infinite_iff_natLeq_ordLeq) - also from inf have "|Field r| =o r" by (simp add: card_of_unique ordIso_symmetric) - finally show ?thesis . -qed - -lemma cinfinite_not_czero: "cinfinite r \ \ (r =o (czero :: 'a rel))" -unfolding cinfinite_def by (metis czeroE finite.emptyI) - -lemma Cinfinite_Cnotzero: "Cinfinite r \ Cnotzero r" -by (metis cinfinite_not_czero) - -lemma Cinfinite_cong: "\r1 =o r2; Cinfinite r1\ \ Cinfinite r2" -by (metis Card_order_ordIso2 card_of_mono2 card_of_ordLeq_infinite cinfinite_def ordIso_iff_ordLeq) - -lemma cinfinite_mono: "\r1 \o r2; cinfinite r1\ \ cinfinite r2" -by (metis card_of_mono2 card_of_ordLeq_infinite cinfinite_def) - - -subsection {* Binary sum *} - -definition csum (infixr "+c" 65) where - "r1 +c r2 \ |Field r1 <+> Field r2|" - -lemma Field_csum: "Field (r +c s) = Inl ` Field r \ Inr ` Field s" - unfolding csum_def Field_card_of by auto - -lemma Card_order_csum: - "Card_order (r1 +c r2)" -unfolding csum_def by (simp add: card_of_Card_order) - -lemma csum_Cnotzero1: - "Cnotzero r1 \ Cnotzero (r1 +c r2)" -unfolding csum_def -by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE) - -lemma card_order_csum: - assumes "card_order r1" "card_order r2" - shows "card_order (r1 +c r2)" -proof - - have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto - thus ?thesis unfolding csum_def by (auto simp: card_of_card_order_on) -qed - -lemma cinfinite_csum: - "cinfinite r1 \ cinfinite r2 \ cinfinite (r1 +c r2)" -unfolding cinfinite_def csum_def by (auto simp: Field_card_of) - -lemma Cinfinite_csum1: - "Cinfinite r1 \ Cinfinite (r1 +c r2)" -unfolding cinfinite_def csum_def by (metis Field_card_of card_of_Card_order finite_Plus_iff) - -lemma Cinfinite_csum: - "Cinfinite r1 \ Cinfinite r2 \ Cinfinite (r1 +c r2)" -unfolding cinfinite_def csum_def by (metis Field_card_of card_of_Card_order finite_Plus_iff) - -lemma Cinfinite_csum_strong: - "\Cinfinite r1; Cinfinite r2\ \ Cinfinite (r1 +c r2)" -by (metis Cinfinite_csum) - -lemma csum_cong: "\p1 =o r1; p2 =o r2\ \ p1 +c p2 =o r1 +c r2" -by (simp only: csum_def ordIso_Plus_cong) - -lemma csum_cong1: "p1 =o r1 \ p1 +c q =o r1 +c q" -by (simp only: csum_def ordIso_Plus_cong1) - -lemma csum_cong2: "p2 =o r2 \ q +c p2 =o q +c r2" -by (simp only: csum_def ordIso_Plus_cong2) - -lemma csum_mono: "\p1 \o r1; p2 \o r2\ \ p1 +c p2 \o r1 +c r2" -by (simp only: csum_def ordLeq_Plus_mono) - -lemma csum_mono1: "p1 \o r1 \ p1 +c q \o r1 +c q" -by (simp only: csum_def ordLeq_Plus_mono1) - -lemma csum_mono2: "p2 \o r2 \ q +c p2 \o q +c r2" -by (simp only: csum_def ordLeq_Plus_mono2) - -lemma ordLeq_csum1: "Card_order p1 \ p1 \o p1 +c p2" -by (simp only: csum_def Card_order_Plus1) - -lemma ordLeq_csum2: "Card_order p2 \ p2 \o p1 +c p2" -by (simp only: csum_def Card_order_Plus2) - -lemma csum_com: "p1 +c p2 =o p2 +c p1" -by (simp only: csum_def card_of_Plus_commute) - -lemma csum_assoc: "(p1 +c p2) +c p3 =o p1 +c p2 +c p3" -by (simp only: csum_def Field_card_of card_of_Plus_assoc) - -lemma Cfinite_csum: "\Cfinite r; Cfinite s\ \ Cfinite (r +c s)" - unfolding cfinite_def csum_def Field_card_of using card_of_card_order_on by simp - -lemma csum_csum: "(r1 +c r2) +c (r3 +c r4) =o (r1 +c r3) +c (r2 +c r4)" -proof - - have "(r1 +c r2) +c (r3 +c r4) =o r1 +c r2 +c (r3 +c r4)" - by (metis csum_assoc) - also have "r1 +c r2 +c (r3 +c r4) =o r1 +c (r2 +c r3) +c r4" - by (metis csum_assoc csum_cong2 ordIso_symmetric) - also have "r1 +c (r2 +c r3) +c r4 =o r1 +c (r3 +c r2) +c r4" - by (metis csum_com csum_cong1 csum_cong2) - also have "r1 +c (r3 +c r2) +c r4 =o r1 +c r3 +c r2 +c r4" - by (metis csum_assoc csum_cong2 ordIso_symmetric) - also have "r1 +c r3 +c r2 +c r4 =o (r1 +c r3) +c (r2 +c r4)" - by (metis csum_assoc ordIso_symmetric) - finally show ?thesis . -qed - -lemma Plus_csum: "|A <+> B| =o |A| +c |B|" -by (simp only: csum_def Field_card_of card_of_refl) - -lemma Un_csum: "|A \ B| \o |A| +c |B|" -using ordLeq_ordIso_trans[OF card_of_Un_Plus_ordLeq Plus_csum] by blast - - -subsection {* One *} - -definition cone where - "cone = card_of {()}" - -lemma Card_order_cone: "Card_order cone" -unfolding cone_def by (rule card_of_Card_order) - -lemma Cfinite_cone: "Cfinite cone" - unfolding cfinite_def by (simp add: Card_order_cone) - -lemma cone_not_czero: "\ (cone =o czero)" -unfolding czero_def cone_def by (metis empty_not_insert card_of_empty3[of "{()}"] ordIso_iff_ordLeq) - -lemma cone_ordLeq_Cnotzero: "Cnotzero r \ cone \o r" -unfolding cone_def by (metis Card_order_singl_ordLeq czeroI) - - -subsection{* Two *} - -definition ctwo where - "ctwo = |UNIV :: bool set|" - -lemma Card_order_ctwo: "Card_order ctwo" -unfolding ctwo_def by (rule card_of_Card_order) - -lemma ctwo_not_czero: "\ (ctwo =o czero)" -using card_of_empty3[of "UNIV :: bool set"] ordIso_iff_ordLeq -unfolding czero_def ctwo_def by (metis UNIV_not_empty) - -lemma ctwo_Cnotzero: "Cnotzero ctwo" -by (simp add: ctwo_not_czero Card_order_ctwo) - - -subsection {* Family sum *} - -definition Csum where - "Csum r rs \ |SIGMA i : Field r. Field (rs i)|" - -(* Similar setup to the one for SIGMA from theory Big_Operators: *) -syntax "_Csum" :: - "pttrn => ('a * 'a) set => 'b * 'b set => (('a * 'b) * ('a * 'b)) set" - ("(3CSUM _:_. _)" [0, 51, 10] 10) - -translations - "CSUM i:r. rs" == "CONST Csum r (%i. rs)" - -lemma SIGMA_CSUM: "|SIGMA i : I. As i| = (CSUM i : |I|. |As i| )" -by (auto simp: Csum_def Field_card_of) - -(* NB: Always, under the cardinal operator, -operations on sets are reduced automatically to operations on cardinals. -This should make cardinal reasoning more direct and natural. *) - - -subsection {* Product *} - -definition cprod (infixr "*c" 80) where - "r1 *c r2 = |Field r1 <*> Field r2|" - -lemma card_order_cprod: - assumes "card_order r1" "card_order r2" - shows "card_order (r1 *c r2)" -proof - - have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto - thus ?thesis by (auto simp: cprod_def card_of_card_order_on) -qed - -lemma Card_order_cprod: "Card_order (r1 *c r2)" -by (simp only: cprod_def Field_card_of card_of_card_order_on) - -lemma cprod_mono1: "p1 \o r1 \ p1 *c q \o r1 *c q" -by (simp only: cprod_def ordLeq_Times_mono1) - -lemma cprod_mono2: "p2 \o r2 \ q *c p2 \o q *c r2" -by (simp only: cprod_def ordLeq_Times_mono2) - -lemma ordLeq_cprod2: "\Cnotzero p1; Card_order p2\ \ p2 \o p1 *c p2" -unfolding cprod_def by (metis Card_order_Times2 czeroI) - -lemma cinfinite_cprod: "\cinfinite r1; cinfinite r2\ \ cinfinite (r1 *c r2)" -by (simp add: cinfinite_def cprod_def Field_card_of infinite_cartesian_product) - -lemma cinfinite_cprod2: "\Cnotzero r1; Cinfinite r2\ \ cinfinite (r1 *c r2)" -by (metis cinfinite_mono ordLeq_cprod2) - -lemma Cinfinite_cprod2: "\Cnotzero r1; Cinfinite r2\ \ Cinfinite (r1 *c r2)" -by (blast intro: cinfinite_cprod2 Card_order_cprod) - -lemma cprod_com: "p1 *c p2 =o p2 *c p1" -by (simp only: cprod_def card_of_Times_commute) - -lemma card_of_Csum_Times: - "\i \ I. |A i| \o |B| \ (CSUM i : |I|. |A i| ) \o |I| *c |B|" -by (simp only: Csum_def cprod_def Field_card_of card_of_Sigma_Times) - -lemma card_of_Csum_Times': - assumes "Card_order r" "\i \ I. |A i| \o r" - shows "(CSUM i : |I|. |A i| ) \o |I| *c r" -proof - - from assms(1) have *: "r =o |Field r|" by (simp add: card_of_unique) - with assms(2) have "\i \ I. |A i| \o |Field r|" by (blast intro: ordLeq_ordIso_trans) - hence "(CSUM i : |I|. |A i| ) \o |I| *c |Field r|" by (simp only: card_of_Csum_Times) - also from * have "|I| *c |Field r| \o |I| *c r" - by (simp only: Field_card_of card_of_refl cprod_def ordIso_imp_ordLeq) - finally show ?thesis . -qed - -lemma cprod_csum_distrib1: "r1 *c r2 +c r1 *c r3 =o r1 *c (r2 +c r3)" -unfolding csum_def cprod_def by (simp add: Field_card_of card_of_Times_Plus_distrib ordIso_symmetric) - -lemma csum_absorb2': "\Card_order r2; r1 \o r2; cinfinite r1 \ cinfinite r2\ \ r1 +c r2 =o r2" -unfolding csum_def by (metis Card_order_Plus_infinite cinfinite_def cinfinite_mono) - -lemma csum_absorb1': - assumes card: "Card_order r2" - and r12: "r1 \o r2" and cr12: "cinfinite r1 \ cinfinite r2" - shows "r2 +c r1 =o r2" -by (rule ordIso_transitive, rule csum_com, rule csum_absorb2', (simp only: assms)+) - -lemma csum_absorb1: "\Cinfinite r2; r1 \o r2\ \ r2 +c r1 =o r2" -by (rule csum_absorb1') auto - - -subsection {* Exponentiation *} - -definition cexp (infixr "^c" 90) where - "r1 ^c r2 \ |Func (Field r2) (Field r1)|" - -lemma Card_order_cexp: "Card_order (r1 ^c r2)" -unfolding cexp_def by (rule card_of_Card_order) - -lemma cexp_mono': - assumes 1: "p1 \o r1" and 2: "p2 \o r2" - and n: "Field p2 = {} \ Field r2 = {}" - shows "p1 ^c p2 \o r1 ^c r2" -proof(cases "Field p1 = {}") - case True - hence "|Field |Func (Field p2) (Field p1)|| \o cone" - unfolding cone_def Field_card_of - by (cases "Field p2 = {}", auto intro: card_of_ordLeqI2 simp: Func_empty) - (metis Func_is_emp card_of_empty ex_in_conv) - hence "|Func (Field p2) (Field p1)| \o cone" by (simp add: Field_card_of cexp_def) - hence "p1 ^c p2 \o cone" unfolding cexp_def . - thus ?thesis - proof (cases "Field p2 = {}") - case True - with n have "Field r2 = {}" . - hence "cone \o r1 ^c r2" unfolding cone_def cexp_def Func_def by (auto intro: card_of_ordLeqI) - thus ?thesis using `p1 ^c p2 \o cone` ordLeq_transitive by auto - next - case False with True have "|Field (p1 ^c p2)| =o czero" - unfolding card_of_ordIso_czero_iff_empty cexp_def Field_card_of Func_def by auto - thus ?thesis unfolding cexp_def card_of_ordIso_czero_iff_empty Field_card_of - by (simp add: card_of_empty) - qed -next - case False - have 1: "|Field p1| \o |Field r1|" and 2: "|Field p2| \o |Field r2|" - using 1 2 by (auto simp: card_of_mono2) - obtain f1 where f1: "f1 ` Field r1 = Field p1" - using 1 unfolding card_of_ordLeq2[OF False, symmetric] by auto - obtain f2 where f2: "inj_on f2 (Field p2)" "f2 ` Field p2 \ Field r2" - using 2 unfolding card_of_ordLeq[symmetric] by blast - have 0: "Func_map (Field p2) f1 f2 ` (Field (r1 ^c r2)) = Field (p1 ^c p2)" - unfolding cexp_def Field_card_of using Func_map_surj[OF f1 f2 n, symmetric] . - have 00: "Field (p1 ^c p2) \ {}" unfolding cexp_def Field_card_of Func_is_emp - using False by simp - show ?thesis - using 0 card_of_ordLeq2[OF 00] unfolding cexp_def Field_card_of by blast -qed - -lemma cexp_mono: - assumes 1: "p1 \o r1" and 2: "p2 \o r2" - and n: "p2 =o czero \ r2 =o czero" and card: "Card_order p2" - shows "p1 ^c p2 \o r1 ^c r2" - by (metis (full_types) "1" "2" card cexp_mono' czeroE czeroI n) - -lemma cexp_mono1: - assumes 1: "p1 \o r1" and q: "Card_order q" - shows "p1 ^c q \o r1 ^c q" -using ordLeq_refl[OF q] by (rule cexp_mono[OF 1]) (auto simp: q) - -lemma cexp_mono2': - assumes 2: "p2 \o r2" and q: "Card_order q" - and n: "Field p2 = {} \ Field r2 = {}" - shows "q ^c p2 \o q ^c r2" -using ordLeq_refl[OF q] by (rule cexp_mono'[OF _ 2 n]) auto - -lemma cexp_mono2: - assumes 2: "p2 \o r2" and q: "Card_order q" - and n: "p2 =o czero \ r2 =o czero" and card: "Card_order p2" - shows "q ^c p2 \o q ^c r2" -using ordLeq_refl[OF q] by (rule cexp_mono[OF _ 2 n card]) auto - -lemma cexp_mono2_Cnotzero: - assumes "p2 \o r2" "Card_order q" "Cnotzero p2" - shows "q ^c p2 \o q ^c r2" -by (metis assms cexp_mono2' czeroI) - -lemma cexp_cong: - assumes 1: "p1 =o r1" and 2: "p2 =o r2" - and Cr: "Card_order r2" - and Cp: "Card_order p2" - shows "p1 ^c p2 =o r1 ^c r2" -proof - - obtain f where "bij_betw f (Field p2) (Field r2)" - using 2 card_of_ordIso[of "Field p2" "Field r2"] card_of_cong by auto - hence 0: "Field p2 = {} \ Field r2 = {}" unfolding bij_betw_def by auto - have r: "p2 =o czero \ r2 =o czero" - and p: "r2 =o czero \ p2 =o czero" - using 0 Cr Cp czeroE czeroI by auto - show ?thesis using 0 1 2 unfolding ordIso_iff_ordLeq - using r p cexp_mono[OF _ _ _ Cp] cexp_mono[OF _ _ _ Cr] by blast -qed - -lemma cexp_cong1: - assumes 1: "p1 =o r1" and q: "Card_order q" - shows "p1 ^c q =o r1 ^c q" -by (rule cexp_cong[OF 1 _ q q]) (rule ordIso_refl[OF q]) - -lemma cexp_cong2: - assumes 2: "p2 =o r2" and q: "Card_order q" and p: "Card_order p2" - shows "q ^c p2 =o q ^c r2" -by (rule cexp_cong[OF _ 2]) (auto simp only: ordIso_refl Card_order_ordIso2[OF p 2] q p) - -lemma cexp_cone: - assumes "Card_order r" - shows "r ^c cone =o r" -proof - - have "r ^c cone =o |Field r|" - unfolding cexp_def cone_def Field_card_of Func_empty - card_of_ordIso[symmetric] bij_betw_def Func_def inj_on_def image_def - by (rule exI[of _ "\f. f ()"]) auto - also have "|Field r| =o r" by (rule card_of_Field_ordIso[OF assms]) - finally show ?thesis . -qed - -lemma cexp_cprod: - assumes r1: "Card_order r1" - shows "(r1 ^c r2) ^c r3 =o r1 ^c (r2 *c r3)" (is "?L =o ?R") -proof - - have "?L =o r1 ^c (r3 *c r2)" - unfolding cprod_def cexp_def Field_card_of - using card_of_Func_Times by(rule ordIso_symmetric) - also have "r1 ^c (r3 *c r2) =o ?R" - apply(rule cexp_cong2) using cprod_com r1 by (auto simp: Card_order_cprod) - finally show ?thesis . -qed - -lemma cprod_infinite1': "\Cinfinite r; Cnotzero p; p \o r\ \ r *c p =o r" -unfolding cinfinite_def cprod_def -by (rule Card_order_Times_infinite[THEN conjunct1]) (blast intro: czeroI)+ - -lemma cexp_cprod_ordLeq: - assumes r1: "Card_order r1" and r2: "Cinfinite r2" - and r3: "Cnotzero r3" "r3 \o r2" - shows "(r1 ^c r2) ^c r3 =o r1 ^c r2" (is "?L =o ?R") -proof- - have "?L =o r1 ^c (r2 *c r3)" using cexp_cprod[OF r1] . - also have "r1 ^c (r2 *c r3) =o ?R" - apply(rule cexp_cong2) - apply(rule cprod_infinite1'[OF r2 r3]) using r1 r2 by (fastforce simp: Card_order_cprod)+ - finally show ?thesis . -qed - -lemma Cnotzero_UNIV: "Cnotzero |UNIV|" -by (auto simp: card_of_Card_order card_of_ordIso_czero_iff_empty) - -lemma ordLess_ctwo_cexp: - assumes "Card_order r" - shows "r o q ^c r" -proof (cases "q =o (czero :: 'a rel)") - case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans) -next - case False - thus ?thesis - apply - - apply (rule ordIso_ordLeq_trans) - apply (rule ordIso_symmetric) - apply (rule cexp_cone) - apply (rule assms(2)) - apply (rule cexp_mono2) - apply (rule cone_ordLeq_Cnotzero) - apply (rule assms(1)) - apply (rule assms(2)) - apply (rule notE) - apply (rule cone_not_czero) - apply assumption - apply (rule Card_order_cone) - done -qed - -lemma ordLeq_cexp2: - assumes "ctwo \o q" "Card_order r" - shows "r \o q ^c r" -proof (cases "r =o (czero :: 'a rel)") - case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans) -next - case False thus ?thesis - apply - - apply (rule ordLess_imp_ordLeq) - apply (rule ordLess_ordLeq_trans) - apply (rule ordLess_ctwo_cexp) - apply (rule assms(2)) - apply (rule cexp_mono1) - apply (rule assms(1)) - apply (rule assms(2)) - done -qed - -lemma cinfinite_cexp: "\ctwo \o q; Cinfinite r\ \ cinfinite (q ^c r)" -by (metis assms cinfinite_mono ordLeq_cexp2) - -lemma Cinfinite_cexp: - "\ctwo \o q; Cinfinite r\ \ Cinfinite (q ^c r)" -by (simp add: cinfinite_cexp Card_order_cexp) - -lemma ctwo_ordLess_natLeq: "ctwo ctwo o r" -by (rule ordLess_imp_ordLeq[OF ctwo_ordLess_Cinfinite[OF assms]]) - -lemma Un_Cinfinite_bound: "\|A| \o r; |B| \o r; Cinfinite r\ \ |A \ B| \o r" -by (auto simp add: cinfinite_def card_of_Un_ordLeq_infinite_Field) - -lemma UNION_Cinfinite_bound: "\|I| \o r; \i \ I. |A i| \o r; Cinfinite r\ \ |\i \ I. A i| \o r" -by (auto simp add: card_of_UNION_ordLeq_infinite_Field cinfinite_def) - -lemma csum_cinfinite_bound: - assumes "p \o r" "q \o r" "Card_order p" "Card_order q" "Cinfinite r" - shows "p +c q \o r" -proof - - from assms(1-4) have "|Field p| \o r" "|Field q| \o r" - unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+ - with assms show ?thesis unfolding cinfinite_def csum_def - by (blast intro: card_of_Plus_ordLeq_infinite_Field) -qed - -lemma cprod_cinfinite_bound: - assumes "p \o r" "q \o r" "Card_order p" "Card_order q" "Cinfinite r" - shows "p *c q \o r" -proof - - from assms(1-4) have "|Field p| \o r" "|Field q| \o r" - unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+ - with assms show ?thesis unfolding cinfinite_def cprod_def - by (blast intro: card_of_Times_ordLeq_infinite_Field) -qed - -lemma cprod_csum_cexp: - "r1 *c r2 \o (r1 +c r2) ^c ctwo" -unfolding cprod_def csum_def cexp_def ctwo_def Field_card_of -proof - - let ?f = "\(a, b). %x. if x then Inl a else Inr b" - have "inj_on ?f (Field r1 \ Field r2)" (is "inj_on _ ?LHS") - by (auto simp: inj_on_def fun_eq_iff split: bool.split) - moreover - have "?f ` ?LHS \ Func (UNIV :: bool set) (Field r1 <+> Field r2)" (is "_ \ ?RHS") - by (auto simp: Func_def) - ultimately show "|?LHS| \o |?RHS|" using card_of_ordLeq by blast -qed - -lemma Cfinite_cprod_Cinfinite: "\Cfinite r; Cinfinite s\ \ r *c s \o s" -by (intro cprod_cinfinite_bound) - (auto intro: ordLeq_refl ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite]) - -lemma cprod_cexp: "(r *c s) ^c t =o r ^c t *c s ^c t" - unfolding cprod_def cexp_def Field_card_of by (rule Func_Times_Range) - -lemma cprod_cexp_csum_cexp_Cinfinite: - assumes t: "Cinfinite t" - shows "(r *c s) ^c t \o (r +c s) ^c t" -proof - - have "(r *c s) ^c t \o ((r +c s) ^c ctwo) ^c t" - by (rule cexp_mono1[OF cprod_csum_cexp conjunct2[OF t]]) - also have "((r +c s) ^c ctwo) ^c t =o (r +c s) ^c (ctwo *c t)" - by (rule cexp_cprod[OF Card_order_csum]) - also have "(r +c s) ^c (ctwo *c t) =o (r +c s) ^c (t *c ctwo)" - by (rule cexp_cong2[OF cprod_com Card_order_csum Card_order_cprod]) - also have "(r +c s) ^c (t *c ctwo) =o ((r +c s) ^c t) ^c ctwo" - by (rule ordIso_symmetric[OF cexp_cprod[OF Card_order_csum]]) - also have "((r +c s) ^c t) ^c ctwo =o (r +c s) ^c t" - by (rule cexp_cprod_ordLeq[OF Card_order_csum t ctwo_Cnotzero ctwo_ordLeq_Cinfinite[OF t]]) - finally show ?thesis . -qed - -lemma Cfinite_cexp_Cinfinite: - assumes s: "Cfinite s" and t: "Cinfinite t" - shows "s ^c t \o ctwo ^c t" -proof (cases "s \o ctwo") - case True thus ?thesis using t by (blast intro: cexp_mono1) -next - case False - hence "ctwo \o s" by (metis card_order_on_well_order_on ctwo_Cnotzero ordLeq_total s) - hence "Cnotzero s" by (metis Cnotzero_mono ctwo_Cnotzero s) - hence st: "Cnotzero (s *c t)" by (metis Cinfinite_cprod2 cinfinite_not_czero t) - have "s ^c t \o (ctwo ^c s) ^c t" - using assms by (blast intro: cexp_mono1 ordLess_imp_ordLeq[OF ordLess_ctwo_cexp]) - also have "(ctwo ^c s) ^c t =o ctwo ^c (s *c t)" - by (blast intro: Card_order_ctwo cexp_cprod) - also have "ctwo ^c (s *c t) \o ctwo ^c t" - using assms st by (intro cexp_mono2_Cnotzero Cfinite_cprod_Cinfinite Card_order_ctwo) - finally show ?thesis . -qed - -lemma csum_Cfinite_cexp_Cinfinite: - assumes r: "Card_order r" and s: "Cfinite s" and t: "Cinfinite t" - shows "(r +c s) ^c t \o (r +c ctwo) ^c t" -proof (cases "Cinfinite r") - case True - hence "r +c s =o r" by (intro csum_absorb1 ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite] s) - hence "(r +c s) ^c t =o r ^c t" using t by (blast intro: cexp_cong1) - also have "r ^c t \o (r +c ctwo) ^c t" using t by (blast intro: cexp_mono1 ordLeq_csum1 r) - finally show ?thesis . -next - case False - with r have "Cfinite r" unfolding cinfinite_def cfinite_def by auto - hence "Cfinite (r +c s)" by (intro Cfinite_csum s) - hence "(r +c s) ^c t \o ctwo ^c t" by (intro Cfinite_cexp_Cinfinite t) - also have "ctwo ^c t \o (r +c ctwo) ^c t" using t - by (blast intro: cexp_mono1 ordLeq_csum2 Card_order_ctwo) - finally show ?thesis . -qed - -lemma card_order_cexp: - assumes "card_order r1" "card_order r2" - shows "card_order (r1 ^c r2)" -proof - - have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto - thus ?thesis unfolding cexp_def Func_def by (simp add: card_of_card_order_on) -qed - -lemma Cinfinite_ordLess_cexp: - assumes r: "Cinfinite r" - shows "r o r ^c r" - by (rule cexp_mono1[OF ctwo_ordLeq_Cinfinite]) (auto simp: r ctwo_not_czero Card_order_ctwo) - finally show ?thesis . -qed - -lemma infinite_ordLeq_cexp: - assumes "Cinfinite r" - shows "r \o r ^c r" -by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]]) - -(* cardSuc *) - -lemma Cinfinite_cardSuc: "Cinfinite r \ Cinfinite (cardSuc r)" -by (simp add: cinfinite_def cardSuc_Card_order cardSuc_finite) - -lemma cardSuc_UNION_Cinfinite: - assumes "Cinfinite r" "relChain (cardSuc r) As" "B \ (UN i : Field (cardSuc r). As i)" "|B| <=o r" - shows "EX i : Field (cardSuc r). B \ As i" -using cardSuc_UNION assms unfolding cinfinite_def by blast - -subsection {* Powerset *} - -definition cpow where "cpow r = |Pow (Field r)|" - -lemma card_order_cpow: "card_order r \ card_order (cpow r)" -by (simp only: cpow_def Field_card_order Pow_UNIV card_of_card_order_on) - -lemma cpow_greater_eq: "Card_order r \ r \o cpow r" -by (rule ordLess_imp_ordLeq) (simp only: cpow_def Card_order_Pow) - -subsection {* Lists *} - -definition clists where "clists r = |lists (Field r)|" - -end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Cardinal_Order_Relation.thy --- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* Cardinal-Order Relations *} theory Cardinal_Order_Relation -imports Cardinal_Order_Relation_LFP Constructions_on_Wellorders +imports Cardinal_Order_Relation_FP Constructions_on_Wellorders begin declare @@ -136,15 +136,7 @@ subsection {* Cardinals versus set operations on arbitrary sets *} -lemma infinite_Pow: -assumes "infinite A" -shows "infinite (Pow A)" -proof- - have "|A| \o |Pow A|" by (metis card_of_Pow ordLess_imp_ordLeq) - thus ?thesis by (metis assms finite_Pow_iff) -qed - -corollary card_of_set_type[simp]: "|UNIV::'a set| o"} (which is the same as being {\em minimal} w.r.t. the +strict order-embedding relation, @{text " 'a rel \ bool" +where +"card_order_on A r \ well_order_on A r \ (\r'. well_order_on A r' \ r \o r')" + + +abbreviation "Card_order r \ card_order_on (Field r) r" +abbreviation "card_order r \ card_order_on UNIV r" + + +lemma card_order_on_well_order_on: +assumes "card_order_on A r" +shows "well_order_on A r" +using assms unfolding card_order_on_def by simp + + +lemma card_order_on_Card_order: +"card_order_on A r \ A = Field r \ Card_order r" +unfolding card_order_on_def using rel.well_order_on_Field by blast + + +text{* The existence of a cardinal relation on any given set (which will mean +that any set has a cardinal) follows from two facts: +\begin{itemize} +\item Zermelo's theorem (proved in @{text "Zorn.thy"} as theorem @{text "well_order_on"}), +which states that on any given set there exists a well-order; +\item The well-founded-ness of @{text "r. card_order_on A r" +proof- + obtain R where R_def: "R = {r. well_order_on A r}" by blast + have 1: "R \ {} \ (\r \ R. Well_order r)" + using well_order_on[of A] R_def rel.well_order_on_Well_order by blast + hence "\r \ R. \r' \ R. r \o r'" + using exists_minim_Well_order[of R] by auto + thus ?thesis using R_def unfolding card_order_on_def by auto +qed + + +lemma card_order_on_ordIso: +assumes CO: "card_order_on A r" and CO': "card_order_on A r'" +shows "r =o r'" +using assms unfolding card_order_on_def +using ordIso_iff_ordLeq by blast + + +lemma Card_order_ordIso: +assumes CO: "Card_order r" and ISO: "r' =o r" +shows "Card_order r'" +using ISO unfolding ordIso_def +proof(unfold card_order_on_def, auto) + fix p' assume "well_order_on (Field r') p'" + hence 0: "Well_order p' \ Field p' = Field r'" + using rel.well_order_on_Well_order by blast + obtain f where 1: "iso r' r f" and 2: "Well_order r \ Well_order r'" + using ISO unfolding ordIso_def by auto + hence 3: "inj_on f (Field r') \ f ` (Field r') = Field r" + by (auto simp add: iso_iff embed_inj_on) + let ?p = "dir_image p' f" + have 4: "p' =o ?p \ Well_order ?p" + using 0 2 3 by (auto simp add: dir_image_ordIso Well_order_dir_image) + moreover have "Field ?p = Field r" + using 0 3 by (auto simp add: dir_image_Field2 order_on_defs) + ultimately have "well_order_on (Field r) ?p" by auto + hence "r \o ?p" using CO unfolding card_order_on_def by auto + thus "r' \o p'" + using ISO 4 ordLeq_ordIso_trans ordIso_ordLeq_trans ordIso_symmetric by blast +qed + + +lemma Card_order_ordIso2: +assumes CO: "Card_order r" and ISO: "r =o r'" +shows "Card_order r'" +using assms Card_order_ordIso ordIso_symmetric by blast + + +subsection {* Cardinal of a set *} + + +text{* We define the cardinal of set to be {\em some} cardinal order on that set. +We shall prove that this notion is unique up to order isomorphism, meaning +that order isomorphism shall be the true identity of cardinals. *} + + +definition card_of :: "'a set \ 'a rel" ("|_|" ) +where "card_of A = (SOME r. card_order_on A r)" + + +lemma card_of_card_order_on: "card_order_on A |A|" +unfolding card_of_def by (auto simp add: card_order_on someI_ex) + + +lemma card_of_well_order_on: "well_order_on A |A|" +using card_of_card_order_on card_order_on_def by blast + + +lemma Field_card_of: "Field |A| = A" +using card_of_card_order_on[of A] unfolding card_order_on_def +using rel.well_order_on_Field by blast + + +lemma card_of_Card_order: "Card_order |A|" +by (simp only: card_of_card_order_on Field_card_of) + + +corollary ordIso_card_of_imp_Card_order: +"r =o |A| \ Card_order r" +using card_of_Card_order Card_order_ordIso by blast + + +lemma card_of_Well_order: "Well_order |A|" +using card_of_Card_order unfolding card_order_on_def by auto + + +lemma card_of_refl: "|A| =o |A|" +using card_of_Well_order ordIso_reflexive by blast + + +lemma card_of_least: "well_order_on A r \ |A| \o r" +using card_of_card_order_on unfolding card_order_on_def by blast + + +lemma card_of_ordIso: +"(\f. bij_betw f A B) = ( |A| =o |B| )" +proof(auto) + fix f assume *: "bij_betw f A B" + then obtain r where "well_order_on B r \ |A| =o r" + using Well_order_iso_copy card_of_well_order_on by blast + hence "|B| \o |A|" using card_of_least + ordLeq_ordIso_trans ordIso_symmetric by blast + moreover + {let ?g = "inv_into A f" + have "bij_betw ?g B A" using * bij_betw_inv_into by blast + then obtain r where "well_order_on A r \ |B| =o r" + using Well_order_iso_copy card_of_well_order_on by blast + hence "|A| \o |B|" using card_of_least + ordLeq_ordIso_trans ordIso_symmetric by blast + } + ultimately show "|A| =o |B|" using ordIso_iff_ordLeq by blast +next + assume "|A| =o |B|" + then obtain f where "iso ( |A| ) ( |B| ) f" + unfolding ordIso_def by auto + hence "bij_betw f A B" unfolding iso_def Field_card_of by simp + thus "\f. bij_betw f A B" by auto +qed + + +lemma card_of_ordLeq: +"(\f. inj_on f A \ f ` A \ B) = ( |A| \o |B| )" +proof(auto) + fix f assume *: "inj_on f A" and **: "f ` A \ B" + {assume "|B| o |A|" using ordLeq_iff_ordLess_or_ordIso by blast + then obtain g where "embed ( |B| ) ( |A| ) g" + unfolding ordLeq_def by auto + hence 1: "inj_on g B \ g ` B \ A" using embed_inj_on[of "|B|" "|A|" "g"] + card_of_Well_order[of "B"] Field_card_of[of "B"] Field_card_of[of "A"] + embed_Field[of "|B|" "|A|" g] by auto + obtain h where "bij_betw h A B" + using * ** 1 Cantor_Bernstein[of f] by fastforce + hence "|A| =o |B|" using card_of_ordIso by blast + hence "|A| \o |B|" using ordIso_iff_ordLeq by auto + } + thus "|A| \o |B|" using ordLess_or_ordLeq[of "|B|" "|A|"] + by (auto simp: card_of_Well_order) +next + assume *: "|A| \o |B|" + obtain f where "embed ( |A| ) ( |B| ) f" + using * unfolding ordLeq_def by auto + hence "inj_on f A \ f ` A \ B" using embed_inj_on[of "|A|" "|B|" f] + card_of_Well_order[of "A"] Field_card_of[of "A"] Field_card_of[of "B"] + embed_Field[of "|A|" "|B|" f] by auto + thus "\f. inj_on f A \ f ` A \ B" by auto +qed + + +lemma card_of_ordLeq2: +"A \ {} \ (\g. g ` B = A) = ( |A| \o |B| )" +using card_of_ordLeq[of A B] inj_on_iff_surj[of A B] by auto + + +lemma card_of_ordLess: +"(\(\f. inj_on f A \ f ` A \ B)) = ( |B| (\f. inj_on f A \ f ` A \ B)) = (\ |A| \o |B| )" + using card_of_ordLeq by blast + also have "\ = ( |B| {} \ (\(\f. f ` A = B)) = ( |A| a. a \ A \ f a \ B" +shows "|A| \o |B|" +using assms unfolding card_of_ordLeq[symmetric] by auto + + +lemma card_of_unique: +"card_order_on A r \ r =o |A|" +by (simp only: card_order_on_ordIso card_of_card_order_on) + + +lemma card_of_mono1: +"A \ B \ |A| \o |B|" +using inj_on_id[of A] card_of_ordLeq[of A B] by fastforce + + +lemma card_of_mono2: +assumes "r \o r'" +shows "|Field r| \o |Field r'|" +proof- + obtain f where + 1: "well_order_on (Field r) r \ well_order_on (Field r) r \ embed r r' f" + using assms unfolding ordLeq_def + by (auto simp add: rel.well_order_on_Well_order) + hence "inj_on f (Field r) \ f ` (Field r) \ Field r'" + by (auto simp add: embed_inj_on embed_Field) + thus "|Field r| \o |Field r'|" using card_of_ordLeq by blast +qed + + +lemma card_of_cong: "r =o r' \ |Field r| =o |Field r'|" +by (simp add: ordIso_iff_ordLeq card_of_mono2) + + +lemma card_of_Field_ordLess: "Well_order r \ |Field r| \o r" +using card_of_least card_of_well_order_on rel.well_order_on_Well_order by blast + + +lemma card_of_Field_ordIso: +assumes "Card_order r" +shows "|Field r| =o r" +proof- + have "card_order_on (Field r) r" + using assms card_order_on_Card_order by blast + moreover have "card_order_on (Field r) |Field r|" + using card_of_card_order_on by blast + ultimately show ?thesis using card_order_on_ordIso by blast +qed + + +lemma Card_order_iff_ordIso_card_of: +"Card_order r = (r =o |Field r| )" +using ordIso_card_of_imp_Card_order card_of_Field_ordIso ordIso_symmetric by blast + + +lemma Card_order_iff_ordLeq_card_of: +"Card_order r = (r \o |Field r| )" +proof- + have "Card_order r = (r =o |Field r| )" + unfolding Card_order_iff_ordIso_card_of by simp + also have "... = (r \o |Field r| \ |Field r| \o r)" + unfolding ordIso_iff_ordLeq by simp + also have "... = (r \o |Field r| )" + using card_of_Field_ordLess + by (auto simp: card_of_Field_ordLess ordLeq_Well_order_simp) + finally show ?thesis . +qed + + +lemma Card_order_iff_Restr_underS: +assumes "Well_order r" +shows "Card_order r = (\a \ Field r. Restr r (rel.underS r a) o ?r'" + unfolding card_order_on_def by simp + moreover have "Field ?r' = ?A" + using 1 wo_rel.underS_ofilter Field_Restr_ofilter + unfolding wo_rel_def by fastforce + ultimately have "|?A| \o ?r'" by simp + also have "?r' o r" using card_of_least by blast + thus ?thesis using assms ordLeq_ordLess_trans by blast +qed + + +lemma internalize_card_of_ordLeq: +"( |A| \o r) = (\B \ Field r. |A| =o |B| \ |B| \o r)" +proof + assume "|A| \o r" + then obtain p where 1: "Field p \ Field r \ |A| =o p \ p \o r" + using internalize_ordLeq[of "|A|" r] by blast + hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast + hence "|Field p| =o p" using card_of_Field_ordIso by blast + hence "|A| =o |Field p| \ |Field p| \o r" + using 1 ordIso_equivalence ordIso_ordLeq_trans by blast + thus "\B \ Field r. |A| =o |B| \ |B| \o r" using 1 by blast +next + assume "\B \ Field r. |A| =o |B| \ |B| \o r" + thus "|A| \o r" using ordIso_ordLeq_trans by blast +qed + + +lemma internalize_card_of_ordLeq2: +"( |A| \o |C| ) = (\B \ C. |A| =o |B| \ |B| \o |C| )" +using internalize_card_of_ordLeq[of "A" "|C|"] Field_card_of[of C] by auto + + + +subsection {* Cardinals versus set operations on arbitrary sets *} + + +text{* Here we embark in a long journey of simple results showing +that the standard set-theoretic operations are well-behaved w.r.t. the notion of +cardinal -- essentially, this means that they preserve the ``cardinal identity" +@{text "=o"} and are monotonic w.r.t. @{text "\o"}. +*} + + +lemma card_of_empty: "|{}| \o |A|" +using card_of_ordLeq inj_on_id by blast + + +lemma card_of_empty1: +assumes "Well_order r \ Card_order r" +shows "|{}| \o r" +proof- + have "Well_order r" using assms unfolding card_order_on_def by auto + hence "|Field r| <=o r" + using assms card_of_Field_ordLess by blast + moreover have "|{}| \o |Field r|" by (simp add: card_of_empty) + ultimately show ?thesis using ordLeq_transitive by blast +qed + + +corollary Card_order_empty: +"Card_order r \ |{}| \o r" by (simp add: card_of_empty1) + + +lemma card_of_empty2: +assumes LEQ: "|A| =o |{}|" +shows "A = {}" +using assms card_of_ordIso[of A] bij_betw_empty2 by blast + + +lemma card_of_empty3: +assumes LEQ: "|A| \o |{}|" +shows "A = {}" +using assms +by (simp add: ordIso_iff_ordLeq card_of_empty1 card_of_empty2 + ordLeq_Well_order_simp) + + +lemma card_of_empty_ordIso: +"|{}::'a set| =o |{}::'b set|" +using card_of_ordIso unfolding bij_betw_def inj_on_def by blast + + +lemma card_of_image: +"|f ` A| <=o |A|" +proof(cases "A = {}", simp add: card_of_empty) + assume "A ~= {}" + hence "f ` A ~= {}" by auto + thus "|f ` A| \o |A|" + using card_of_ordLeq2[of "f ` A" A] by auto +qed + + +lemma surj_imp_ordLeq: +assumes "B <= f ` A" +shows "|B| <=o |A|" +proof- + have "|B| <=o |f ` A|" using assms card_of_mono1 by auto + thus ?thesis using card_of_image ordLeq_transitive by blast +qed + + +lemma card_of_ordLeqI2: +assumes "B \ f ` A" +shows "|B| \o |A|" +using assms by (metis surj_imp_ordLeq) + + +lemma card_of_singl_ordLeq: +assumes "A \ {}" +shows "|{b}| \o |A|" +proof- + obtain a where *: "a \ A" using assms by auto + let ?h = "\ b'::'b. if b' = b then a else undefined" + have "inj_on ?h {b} \ ?h ` {b} \ A" + using * unfolding inj_on_def by auto + thus ?thesis using card_of_ordLeq by blast +qed + + +corollary Card_order_singl_ordLeq: +"\Card_order r; Field r \ {}\ \ |{b}| \o r" +using card_of_singl_ordLeq[of "Field r" b] + card_of_Field_ordIso[of r] ordLeq_ordIso_trans by blast + + +lemma card_of_Pow: "|A| r o |Pow A|" by (metis card_of_Pow ordLess_imp_ordLeq) + thus ?thesis by (metis assms finite_Pow_iff) +qed + + +lemma card_of_Plus1: "|A| \o |A <+> B|" +proof- + have "Inl ` A \ A <+> B" by auto + thus ?thesis using inj_Inl[of A] card_of_ordLeq by blast +qed + + +corollary Card_order_Plus1: +"Card_order r \ r \o |(Field r) <+> B|" +using card_of_Plus1 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast + + +lemma card_of_Plus2: "|B| \o |A <+> B|" +proof- + have "Inr ` B \ A <+> B" by auto + thus ?thesis using inj_Inr[of B] card_of_ordLeq by blast +qed + + +corollary Card_order_Plus2: +"Card_order r \ r \o |A <+> (Field r)|" +using card_of_Plus2 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast + + +lemma card_of_Plus_empty1: "|A| =o |A <+> {}|" +proof- + have "bij_betw Inl A (A <+> {})" unfolding bij_betw_def inj_on_def by auto + thus ?thesis using card_of_ordIso by auto +qed + + +lemma card_of_Plus_empty2: "|A| =o |{} <+> A|" +proof- + have "bij_betw Inr A ({} <+> A)" unfolding bij_betw_def inj_on_def by auto + thus ?thesis using card_of_ordIso by auto +qed + + +lemma card_of_Plus_commute: "|A <+> B| =o |B <+> A|" +proof- + let ?f = "\(c::'a + 'b). case c of Inl a \ Inr a + | Inr b \ Inl b" + have "bij_betw ?f (A <+> B) (B <+> A)" + unfolding bij_betw_def inj_on_def by force + thus ?thesis using card_of_ordIso by blast +qed + + +lemma card_of_Plus_assoc: +fixes A :: "'a set" and B :: "'b set" and C :: "'c set" +shows "|(A <+> B) <+> C| =o |A <+> B <+> C|" +proof - + def f \ "\(k::('a + 'b) + 'c). + case k of Inl ab \ (case ab of Inl a \ Inl a + |Inr b \ Inr (Inl b)) + |Inr c \ Inr (Inr c)" + have "A <+> B <+> C \ f ` ((A <+> B) <+> C)" + proof + fix x assume x: "x \ A <+> B <+> C" + show "x \ f ` ((A <+> B) <+> C)" + proof(cases x) + case (Inl a) + hence "a \ A" "x = f (Inl (Inl a))" + using x unfolding f_def by auto + thus ?thesis by auto + next + case (Inr bc) note 1 = Inr show ?thesis + proof(cases bc) + case (Inl b) + hence "b \ B" "x = f (Inl (Inr b))" + using x 1 unfolding f_def by auto + thus ?thesis by auto + next + case (Inr c) + hence "c \ C" "x = f (Inr c)" + using x 1 unfolding f_def by auto + thus ?thesis by auto + qed + qed + qed + hence "bij_betw f ((A <+> B) <+> C) (A <+> B <+> C)" + unfolding bij_betw_def inj_on_def f_def by force + thus ?thesis using card_of_ordIso by blast +qed + + +lemma card_of_Plus_mono1: +assumes "|A| \o |B|" +shows "|A <+> C| \o |B <+> C|" +proof- + obtain f where 1: "inj_on f A \ f ` A \ B" + using assms card_of_ordLeq[of A] by fastforce + obtain g where g_def: + "g = (\d. case d of Inl a \ Inl(f a) | Inr (c::'c) \ Inr c)" by blast + have "inj_on g (A <+> C) \ g ` (A <+> C) \ (B <+> C)" + proof- + {fix d1 and d2 assume "d1 \ A <+> C \ d2 \ A <+> C" and + "g d1 = g d2" + hence "d1 = d2" using 1 unfolding inj_on_def + by(case_tac d1, case_tac d2, auto simp add: g_def) + } + moreover + {fix d assume "d \ A <+> C" + hence "g d \ B <+> C" using 1 + by(case_tac d, auto simp add: g_def) + } + ultimately show ?thesis unfolding inj_on_def by auto + qed + thus ?thesis using card_of_ordLeq by metis +qed + + +corollary ordLeq_Plus_mono1: +assumes "r \o r'" +shows "|(Field r) <+> C| \o |(Field r') <+> C|" +using assms card_of_mono2 card_of_Plus_mono1 by blast + + +lemma card_of_Plus_mono2: +assumes "|A| \o |B|" +shows "|C <+> A| \o |C <+> B|" +using assms card_of_Plus_mono1[of A B C] + card_of_Plus_commute[of C A] card_of_Plus_commute[of B C] + ordIso_ordLeq_trans[of "|C <+> A|"] ordLeq_ordIso_trans[of "|C <+> A|"] +by blast + + +corollary ordLeq_Plus_mono2: +assumes "r \o r'" +shows "|A <+> (Field r)| \o |A <+> (Field r')|" +using assms card_of_mono2 card_of_Plus_mono2 by blast + + +lemma card_of_Plus_mono: +assumes "|A| \o |B|" and "|C| \o |D|" +shows "|A <+> C| \o |B <+> D|" +using assms card_of_Plus_mono1[of A B C] card_of_Plus_mono2[of C D B] + ordLeq_transitive[of "|A <+> C|"] by blast + + +corollary ordLeq_Plus_mono: +assumes "r \o r'" and "p \o p'" +shows "|(Field r) <+> (Field p)| \o |(Field r') <+> (Field p')|" +using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Plus_mono by blast + + +lemma card_of_Plus_cong1: +assumes "|A| =o |B|" +shows "|A <+> C| =o |B <+> C|" +using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono1) + + +corollary ordIso_Plus_cong1: +assumes "r =o r'" +shows "|(Field r) <+> C| =o |(Field r') <+> C|" +using assms card_of_cong card_of_Plus_cong1 by blast + + +lemma card_of_Plus_cong2: +assumes "|A| =o |B|" +shows "|C <+> A| =o |C <+> B|" +using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono2) + + +corollary ordIso_Plus_cong2: +assumes "r =o r'" +shows "|A <+> (Field r)| =o |A <+> (Field r')|" +using assms card_of_cong card_of_Plus_cong2 by blast + + +lemma card_of_Plus_cong: +assumes "|A| =o |B|" and "|C| =o |D|" +shows "|A <+> C| =o |B <+> D|" +using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono) + + +corollary ordIso_Plus_cong: +assumes "r =o r'" and "p =o p'" +shows "|(Field r) <+> (Field p)| =o |(Field r') <+> (Field p')|" +using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Plus_cong by blast + + +lemma card_of_Un_Plus_ordLeq: +"|A \ B| \o |A <+> B|" +proof- + let ?f = "\ c. if c \ A then Inl c else Inr c" + have "inj_on ?f (A \ B) \ ?f ` (A \ B) \ A <+> B" + unfolding inj_on_def by auto + thus ?thesis using card_of_ordLeq by blast +qed + + +lemma card_of_Times1: +assumes "A \ {}" +shows "|B| \o |B \ A|" +proof(cases "B = {}", simp add: card_of_empty) + assume *: "B \ {}" + have "fst `(B \ A) = B" unfolding image_def using assms by auto + thus ?thesis using inj_on_iff_surj[of B "B \ A"] + card_of_ordLeq[of B "B \ A"] * by blast +qed + + +lemma card_of_Times_commute: "|A \ B| =o |B \ A|" +proof- + let ?f = "\(a::'a,b::'b). (b,a)" + have "bij_betw ?f (A \ B) (B \ A)" + unfolding bij_betw_def inj_on_def by auto + thus ?thesis using card_of_ordIso by blast +qed + + +lemma card_of_Times2: +assumes "A \ {}" shows "|B| \o |A \ B|" +using assms card_of_Times1[of A B] card_of_Times_commute[of B A] + ordLeq_ordIso_trans by blast + + +corollary Card_order_Times1: +"\Card_order r; B \ {}\ \ r \o |(Field r) \ B|" +using card_of_Times1[of B] card_of_Field_ordIso + ordIso_ordLeq_trans ordIso_symmetric by blast + + +corollary Card_order_Times2: +"\Card_order r; A \ {}\ \ r \o |A \ (Field r)|" +using card_of_Times2[of A] card_of_Field_ordIso + ordIso_ordLeq_trans ordIso_symmetric by blast + + +lemma card_of_Times3: "|A| \o |A \ A|" +using card_of_Times1[of A] +by(cases "A = {}", simp add: card_of_empty, blast) + + +lemma card_of_Plus_Times_bool: "|A <+> A| =o |A \ (UNIV::bool set)|" +proof- + let ?f = "\c::'a + 'a. case c of Inl a \ (a,True) + |Inr a \ (a,False)" + have "bij_betw ?f (A <+> A) (A \ (UNIV::bool set))" + proof- + {fix c1 and c2 assume "?f c1 = ?f c2" + hence "c1 = c2" + by(case_tac "c1", case_tac "c2", auto, case_tac "c2", auto) + } + moreover + {fix c assume "c \ A <+> A" + hence "?f c \ A \ (UNIV::bool set)" + by(case_tac c, auto) + } + moreover + {fix a bl assume *: "(a,bl) \ A \ (UNIV::bool set)" + have "(a,bl) \ ?f ` ( A <+> A)" + proof(cases bl) + assume bl hence "?f(Inl a) = (a,bl)" by auto + thus ?thesis using * by force + next + assume "\ bl" hence "?f(Inr a) = (a,bl)" by auto + thus ?thesis using * by force + qed + } + ultimately show ?thesis unfolding bij_betw_def inj_on_def by auto + qed + thus ?thesis using card_of_ordIso by blast +qed + + +lemma card_of_Times_mono1: +assumes "|A| \o |B|" +shows "|A \ C| \o |B \ C|" +proof- + obtain f where 1: "inj_on f A \ f ` A \ B" + using assms card_of_ordLeq[of A] by fastforce + obtain g where g_def: + "g = (\(a,c::'c). (f a,c))" by blast + have "inj_on g (A \ C) \ g ` (A \ C) \ (B \ C)" + using 1 unfolding inj_on_def using g_def by auto + thus ?thesis using card_of_ordLeq by metis +qed + + +corollary ordLeq_Times_mono1: +assumes "r \o r'" +shows "|(Field r) \ C| \o |(Field r') \ C|" +using assms card_of_mono2 card_of_Times_mono1 by blast + + +lemma card_of_Times_mono2: +assumes "|A| \o |B|" +shows "|C \ A| \o |C \ B|" +using assms card_of_Times_mono1[of A B C] + card_of_Times_commute[of C A] card_of_Times_commute[of B C] + ordIso_ordLeq_trans[of "|C \ A|"] ordLeq_ordIso_trans[of "|C \ A|"] +by blast + + +corollary ordLeq_Times_mono2: +assumes "r \o r'" +shows "|A \ (Field r)| \o |A \ (Field r')|" +using assms card_of_mono2 card_of_Times_mono2 by blast + + +lemma card_of_Sigma_mono1: +assumes "\i \ I. |A i| \o |B i|" +shows "|SIGMA i : I. A i| \o |SIGMA i : I. B i|" +proof- + have "\i. i \ I \ (\f. inj_on f (A i) \ f ` (A i) \ B i)" + using assms by (auto simp add: card_of_ordLeq) + with choice[of "\ i f. i \ I \ inj_on f (A i) \ f ` (A i) \ B i"] + obtain F where 1: "\i \ I. inj_on (F i) (A i) \ (F i) ` (A i) \ B i" by metis + obtain g where g_def: "g = (\(i,a::'b). (i,F i a))" by blast + have "inj_on g (Sigma I A) \ g ` (Sigma I A) \ (Sigma I B)" + using 1 unfolding inj_on_def using g_def by force + thus ?thesis using card_of_ordLeq by metis +qed + + +corollary card_of_Sigma_Times: +"\i \ I. |A i| \o |B| \ |SIGMA i : I. A i| \o |I \ B|" +using card_of_Sigma_mono1[of I A "\i. B"] . + + +lemma card_of_UNION_Sigma: +"|\i \ I. A i| \o |SIGMA i : I. A i|" +using Ex_inj_on_UNION_Sigma[of I A] card_of_ordLeq by metis + + +lemma card_of_bool: +assumes "a1 \ a2" +shows "|UNIV::bool set| =o |{a1,a2}|" +proof- + let ?f = "\ bl. case bl of True \ a1 | False \ a2" + have "bij_betw ?f UNIV {a1,a2}" + proof- + {fix bl1 and bl2 assume "?f bl1 = ?f bl2" + hence "bl1 = bl2" using assms by (case_tac bl1, case_tac bl2, auto) + } + moreover + {fix bl have "?f bl \ {a1,a2}" by (case_tac bl, auto) + } + moreover + {fix a assume *: "a \ {a1,a2}" + have "a \ ?f ` UNIV" + proof(cases "a = a1") + assume "a = a1" + hence "?f True = a" by auto thus ?thesis by blast + next + assume "a \ a1" hence "a = a2" using * by auto + hence "?f False = a" by auto thus ?thesis by blast + qed + } + ultimately show ?thesis unfolding bij_betw_def inj_on_def + by (metis image_subsetI order_eq_iff subsetI) + qed + thus ?thesis using card_of_ordIso by blast +qed + + +lemma card_of_Plus_Times_aux: +assumes A2: "a1 \ a2 \ {a1,a2} \ A" and + LEQ: "|A| \o |B|" +shows "|A <+> B| \o |A \ B|" +proof- + have 1: "|UNIV::bool set| \o |A|" + using A2 card_of_mono1[of "{a1,a2}"] card_of_bool[of a1 a2] + ordIso_ordLeq_trans[of "|UNIV::bool set|"] by metis + (* *) + have "|A <+> B| \o |B <+> B|" + using LEQ card_of_Plus_mono1 by blast + moreover have "|B <+> B| =o |B \ (UNIV::bool set)|" + using card_of_Plus_Times_bool by blast + moreover have "|B \ (UNIV::bool set)| \o |B \ A|" + using 1 by (simp add: card_of_Times_mono2) + moreover have " |B \ A| =o |A \ B|" + using card_of_Times_commute by blast + ultimately show "|A <+> B| \o |A \ B|" + using ordLeq_ordIso_trans[of "|A <+> B|" "|B <+> B|" "|B \ (UNIV::bool set)|"] + ordLeq_transitive[of "|A <+> B|" "|B \ (UNIV::bool set)|" "|B \ A|"] + ordLeq_ordIso_trans[of "|A <+> B|" "|B \ A|" "|A \ B|"] + by blast +qed + + +lemma card_of_Plus_Times: +assumes A2: "a1 \ a2 \ {a1,a2} \ A" and + B2: "b1 \ b2 \ {b1,b2} \ B" +shows "|A <+> B| \o |A \ B|" +proof- + {assume "|A| \o |B|" + hence ?thesis using assms by (auto simp add: card_of_Plus_Times_aux) + } + moreover + {assume "|B| \o |A|" + hence "|B <+> A| \o |B \ A|" + using assms by (auto simp add: card_of_Plus_Times_aux) + hence ?thesis + using card_of_Plus_commute card_of_Times_commute + ordIso_ordLeq_trans ordLeq_ordIso_trans by metis + } + ultimately show ?thesis + using card_of_Well_order[of A] card_of_Well_order[of B] + ordLeq_total[of "|A|"] by metis +qed + + +lemma card_of_ordLeq_finite: +assumes "|A| \o |B|" and "finite B" +shows "finite A" +using assms unfolding ordLeq_def +using embed_inj_on[of "|A|" "|B|"] embed_Field[of "|A|" "|B|"] + Field_card_of[of "A"] Field_card_of[of "B"] inj_on_finite[of _ "A" "B"] by fastforce + + +lemma card_of_ordLeq_infinite: +assumes "|A| \o |B|" and "infinite A" +shows "infinite B" +using assms card_of_ordLeq_finite by auto + + +lemma card_of_ordIso_finite: +assumes "|A| =o |B|" +shows "finite A = finite B" +using assms unfolding ordIso_def iso_def[abs_def] +by (auto simp: bij_betw_finite Field_card_of) + + +lemma card_of_ordIso_finite_Field: +assumes "Card_order r" and "r =o |A|" +shows "finite(Field r) = finite A" +using assms card_of_Field_ordIso card_of_ordIso_finite ordIso_equivalence by blast + + +subsection {* Cardinals versus set operations involving infinite sets *} + + +text{* Here we show that, for infinite sets, most set-theoretic constructions +do not increase the cardinality. The cornerstone for this is +theorem @{text "Card_order_Times_same_infinite"}, which states that self-product +does not increase cardinality -- the proof of this fact adapts a standard +set-theoretic argument, as presented, e.g., in the proof of theorem 1.5.11 +at page 47 in \cite{card-book}. Then everything else follows fairly easily. *} + + +lemma infinite_iff_card_of_nat: +"infinite A = ( |UNIV::nat set| \o |A| )" +by (auto simp add: infinite_iff_countable_subset card_of_ordLeq) + + +text{* The next two results correspond to the ZF fact that all infinite cardinals are +limit ordinals: *} + +lemma Card_order_infinite_not_under: +assumes CARD: "Card_order r" and INF: "infinite (Field r)" +shows "\ (\a. Field r = rel.under r a)" +proof(auto) + have 0: "Well_order r \ wo_rel r \ Refl r" + using CARD unfolding wo_rel_def card_order_on_def order_on_defs by auto + fix a assume *: "Field r = rel.under r a" + show False + proof(cases "a \ Field r") + assume Case1: "a \ Field r" + hence "rel.under r a = {}" unfolding Field_def rel.under_def by auto + thus False using INF * by auto + next + let ?r' = "Restr r (rel.underS r a)" + assume Case2: "a \ Field r" + hence 1: "rel.under r a = rel.underS r a \ {a} \ a \ rel.underS r a" + using 0 rel.Refl_under_underS rel.underS_notIn by fastforce + have 2: "wo_rel.ofilter r (rel.underS r a) \ rel.underS r a < Field r" + using 0 wo_rel.underS_ofilter * 1 Case2 by auto + hence "?r' Well_order ?r'" + using 2 0 Field_Restr_ofilter[of r] Well_order_Restr[of r] by blast + ultimately have "|rel.underS r a| f. bij_betw f (rel.under r a) (rel.underS r a)" + using infinite_imp_bij_betw[of "Field r" a] INF * 1 by auto + hence "|rel.under r a| =o |rel.underS r a|" using card_of_ordIso by blast + } + ultimately show False using not_ordLess_ordIso ordIso_symmetric by blast + qed +qed + + +lemma infinite_Card_order_limit: +assumes r: "Card_order r" and "infinite (Field r)" +and a: "a : Field r" +shows "EX b : Field r. a \ b \ (a,b) : r" +proof- + have "Field r \ rel.under r a" + using assms Card_order_infinite_not_under by blast + moreover have "rel.under r a \ Field r" + using rel.under_Field . + ultimately have "rel.under r a < Field r" by blast + then obtain b where 1: "b : Field r \ ~ (b,a) : r" + unfolding rel.under_def by blast + moreover have ba: "b \ a" + using 1 r unfolding card_order_on_def well_order_on_def + linear_order_on_def partial_order_on_def preorder_on_def refl_on_def by auto + ultimately have "(a,b) : r" + using a r unfolding card_order_on_def well_order_on_def linear_order_on_def + total_on_def by blast + thus ?thesis using 1 ba by auto +qed + + +theorem Card_order_Times_same_infinite: +assumes CO: "Card_order r" and INF: "infinite(Field r)" +shows "|Field r \ Field r| \o r" +proof- + obtain phi where phi_def: + "phi = (\r::'a rel. Card_order r \ infinite(Field r) \ + \ |Field r \ Field r| \o r )" by blast + have temp1: "\r. phi r \ Well_order r" + unfolding phi_def card_order_on_def by auto + have Ft: "\(\r. phi r)" + proof + assume "\r. phi r" + hence "{r. phi r} \ {} \ {r. phi r} \ {r. Well_order r}" + using temp1 by auto + then obtain r where 1: "phi r" and 2: "\r'. phi r' \ r \o r'" and + 3: "Card_order r \ Well_order r" + using exists_minim_Well_order[of "{r. phi r}"] temp1 phi_def by blast + let ?A = "Field r" let ?r' = "bsqr r" + have 4: "Well_order ?r' \ Field ?r' = ?A \ ?A \ |?A| =o r" + using 3 bsqr_Well_order Field_bsqr card_of_Field_ordIso by blast + have 5: "Card_order |?A \ ?A| \ Well_order |?A \ ?A|" + using card_of_Card_order card_of_Well_order by blast + (* *) + have "r ?A|" + using 1 3 5 ordLess_or_ordLeq unfolding phi_def by blast + moreover have "|?A \ ?A| \o ?r'" + using card_of_least[of "?A \ ?A"] 4 by auto + ultimately have "r bij_betw f ?A (?A \ ?A)" + unfolding ordLess_def embedS_def[abs_def] + by (auto simp add: Field_bsqr) + let ?B = "f ` ?A" + have "|?A| =o |?B|" + using 3 6 embed_inj_on inj_on_imp_bij_betw card_of_ordIso by blast + hence 8: "r =o |?B|" using 4 ordIso_transitive ordIso_symmetric by blast + (* *) + have "wo_rel.ofilter ?r' ?B" + using 6 embed_Field_ofilter 3 4 by blast + hence "wo_rel.ofilter ?r' ?B \ ?B \ ?A \ ?A \ ?B \ Field ?r'" + using 7 unfolding bij_betw_def using 6 3 embed_inj_on 4 by auto + hence temp2: "wo_rel.ofilter ?r' ?B \ ?B < ?A \ ?A" + using 4 wo_rel_def[of ?r'] wo_rel.ofilter_def[of ?r' ?B] by blast + have "\ (\a. Field r = rel.under r a)" + using 1 unfolding phi_def using Card_order_infinite_not_under[of r] by auto + then obtain A1 where temp3: "wo_rel.ofilter r A1 \ A1 < ?A" and 9: "?B \ A1 \ A1" + using temp2 3 bsqr_ofilter[of r ?B] by blast + hence "|?B| \o |A1 \ A1|" using card_of_mono1 by blast + hence 10: "r \o |A1 \ A1|" using 8 ordIso_ordLeq_trans by blast + let ?r1 = "Restr r A1" + have "?r1 o ?r1" using 3 Well_order_Restr card_of_least by blast + } + ultimately have 11: "|A1| Well_order |A1| \ Card_order |A1|" + using card_of_Card_order[of A1] card_of_Well_order[of A1] + by (simp add: Field_card_of) + moreover have "\ r \o | A1 |" + using temp4 11 3 using not_ordLeq_iff_ordLess by blast + ultimately have "infinite(Field |A1| ) \ Card_order |A1| \ \ r \o | A1 |" + by (simp add: card_of_card_order_on) + hence "|Field |A1| \ Field |A1| | \o |A1|" + using 2 unfolding phi_def by blast + hence "|A1 \ A1 | \o |A1|" using temp4 by auto + hence "r \o |A1|" using 10 ordLeq_transitive by blast + thus False using 11 not_ordLess_ordLeq by auto + qed + thus ?thesis using assms unfolding phi_def by blast +qed + + +corollary card_of_Times_same_infinite: +assumes "infinite A" +shows "|A \ A| =o |A|" +proof- + let ?r = "|A|" + have "Field ?r = A \ Card_order ?r" + using Field_card_of card_of_Card_order[of A] by fastforce + hence "|A \ A| \o |A|" + using Card_order_Times_same_infinite[of ?r] assms by auto + thus ?thesis using card_of_Times3 ordIso_iff_ordLeq by blast +qed + + +lemma card_of_Times_infinite: +assumes INF: "infinite A" and NE: "B \ {}" and LEQ: "|B| \o |A|" +shows "|A \ B| =o |A| \ |B \ A| =o |A|" +proof- + have "|A| \o |A \ B| \ |A| \o |B \ A|" + using assms by (simp add: card_of_Times1 card_of_Times2) + moreover + {have "|A \ B| \o |A \ A| \ |B \ A| \o |A \ A|" + using LEQ card_of_Times_mono1 card_of_Times_mono2 by blast + moreover have "|A \ A| =o |A|" using INF card_of_Times_same_infinite by blast + ultimately have "|A \ B| \o |A| \ |B \ A| \o |A|" + using ordLeq_ordIso_trans[of "|A \ B|"] ordLeq_ordIso_trans[of "|B \ A|"] by auto + } + ultimately show ?thesis by (simp add: ordIso_iff_ordLeq) +qed + + +corollary Card_order_Times_infinite: +assumes INF: "infinite(Field r)" and CARD: "Card_order r" and + NE: "Field p \ {}" and LEQ: "p \o r" +shows "| (Field r) \ (Field p) | =o r \ | (Field p) \ (Field r) | =o r" +proof- + have "|Field r \ Field p| =o |Field r| \ |Field p \ Field r| =o |Field r|" + using assms by (simp add: card_of_Times_infinite card_of_mono2) + thus ?thesis + using assms card_of_Field_ordIso[of r] + ordIso_transitive[of "|Field r \ Field p|"] + ordIso_transitive[of _ "|Field r|"] by blast +qed + + +lemma card_of_Sigma_ordLeq_infinite: +assumes INF: "infinite B" and + LEQ_I: "|I| \o |B|" and LEQ: "\i \ I. |A i| \o |B|" +shows "|SIGMA i : I. A i| \o |B|" +proof(cases "I = {}", simp add: card_of_empty) + assume *: "I \ {}" + have "|SIGMA i : I. A i| \o |I \ B|" + using LEQ card_of_Sigma_Times by blast + moreover have "|I \ B| =o |B|" + using INF * LEQ_I by (auto simp add: card_of_Times_infinite) + ultimately show ?thesis using ordLeq_ordIso_trans by blast +qed + + +lemma card_of_Sigma_ordLeq_infinite_Field: +assumes INF: "infinite (Field r)" and r: "Card_order r" and + LEQ_I: "|I| \o r" and LEQ: "\i \ I. |A i| \o r" +shows "|SIGMA i : I. A i| \o r" +proof- + let ?B = "Field r" + have 1: "r =o |?B| \ |?B| =o r" using r card_of_Field_ordIso + ordIso_symmetric by blast + hence "|I| \o |?B|" "\i \ I. |A i| \o |?B|" + using LEQ_I LEQ ordLeq_ordIso_trans by blast+ + hence "|SIGMA i : I. A i| \o |?B|" using INF LEQ + card_of_Sigma_ordLeq_infinite by blast + thus ?thesis using 1 ordLeq_ordIso_trans by blast +qed + + +lemma card_of_Times_ordLeq_infinite_Field: +"\infinite (Field r); |A| \o r; |B| \o r; Card_order r\ + \ |A <*> B| \o r" +by(simp add: card_of_Sigma_ordLeq_infinite_Field) + + +lemma card_of_Times_infinite_simps: +"\infinite A; B \ {}; |B| \o |A|\ \ |A \ B| =o |A|" +"\infinite A; B \ {}; |B| \o |A|\ \ |A| =o |A \ B|" +"\infinite A; B \ {}; |B| \o |A|\ \ |B \ A| =o |A|" +"\infinite A; B \ {}; |B| \o |A|\ \ |A| =o |B \ A|" +by (auto simp add: card_of_Times_infinite ordIso_symmetric) + + +lemma card_of_UNION_ordLeq_infinite: +assumes INF: "infinite B" and + LEQ_I: "|I| \o |B|" and LEQ: "\i \ I. |A i| \o |B|" +shows "|\ i \ I. A i| \o |B|" +proof(cases "I = {}", simp add: card_of_empty) + assume *: "I \ {}" + have "|\ i \ I. A i| \o |SIGMA i : I. A i|" + using card_of_UNION_Sigma by blast + moreover have "|SIGMA i : I. A i| \o |B|" + using assms card_of_Sigma_ordLeq_infinite by blast + ultimately show ?thesis using ordLeq_transitive by blast +qed + + +corollary card_of_UNION_ordLeq_infinite_Field: +assumes INF: "infinite (Field r)" and r: "Card_order r" and + LEQ_I: "|I| \o r" and LEQ: "\i \ I. |A i| \o r" +shows "|\ i \ I. A i| \o r" +proof- + let ?B = "Field r" + have 1: "r =o |?B| \ |?B| =o r" using r card_of_Field_ordIso + ordIso_symmetric by blast + hence "|I| \o |?B|" "\i \ I. |A i| \o |?B|" + using LEQ_I LEQ ordLeq_ordIso_trans by blast+ + hence "|\ i \ I. A i| \o |?B|" using INF LEQ + card_of_UNION_ordLeq_infinite by blast + thus ?thesis using 1 ordLeq_ordIso_trans by blast +qed + + +lemma card_of_Plus_infinite1: +assumes INF: "infinite A" and LEQ: "|B| \o |A|" +shows "|A <+> B| =o |A|" +proof(cases "B = {}", simp add: card_of_Plus_empty1 card_of_Plus_empty2 ordIso_symmetric) + let ?Inl = "Inl::'a \ 'a + 'b" let ?Inr = "Inr::'b \ 'a + 'b" + assume *: "B \ {}" + then obtain b1 where 1: "b1 \ B" by blast + show ?thesis + proof(cases "B = {b1}") + assume Case1: "B = {b1}" + have 2: "bij_betw ?Inl A ((?Inl ` A))" + unfolding bij_betw_def inj_on_def by auto + hence 3: "infinite (?Inl ` A)" + using INF bij_betw_finite[of ?Inl A] by blast + let ?A' = "?Inl ` A \ {?Inr b1}" + obtain g where "bij_betw g (?Inl ` A) ?A'" + using 3 infinite_imp_bij_betw2[of "?Inl ` A"] by auto + moreover have "?A' = A <+> B" using Case1 by blast + ultimately have "bij_betw g (?Inl ` A) (A <+> B)" by simp + hence "bij_betw (g o ?Inl) A (A <+> B)" + using 2 by (auto simp add: bij_betw_trans) + thus ?thesis using card_of_ordIso ordIso_symmetric by blast + next + assume Case2: "B \ {b1}" + with * 1 obtain b2 where 3: "b1 \ b2 \ {b1,b2} \ B" by fastforce + obtain f where "inj_on f B \ f ` B \ A" + using LEQ card_of_ordLeq[of B] by fastforce + with 3 have "f b1 \ f b2 \ {f b1, f b2} \ A" + unfolding inj_on_def by auto + with 3 have "|A <+> B| \o |A \ B|" + by (auto simp add: card_of_Plus_Times) + moreover have "|A \ B| =o |A|" + using assms * by (simp add: card_of_Times_infinite_simps) + ultimately have "|A <+> B| \o |A|" using ordLeq_ordIso_trans by metis + thus ?thesis using card_of_Plus1 ordIso_iff_ordLeq by blast + qed +qed + + +lemma card_of_Plus_infinite2: +assumes INF: "infinite A" and LEQ: "|B| \o |A|" +shows "|B <+> A| =o |A|" +using assms card_of_Plus_commute card_of_Plus_infinite1 +ordIso_equivalence by blast + + +lemma card_of_Plus_infinite: +assumes INF: "infinite A" and LEQ: "|B| \o |A|" +shows "|A <+> B| =o |A| \ |B <+> A| =o |A|" +using assms by (auto simp: card_of_Plus_infinite1 card_of_Plus_infinite2) + + +corollary Card_order_Plus_infinite: +assumes INF: "infinite(Field r)" and CARD: "Card_order r" and + LEQ: "p \o r" +shows "| (Field r) <+> (Field p) | =o r \ | (Field p) <+> (Field r) | =o r" +proof- + have "| Field r <+> Field p | =o | Field r | \ + | Field p <+> Field r | =o | Field r |" + using assms by (simp add: card_of_Plus_infinite card_of_mono2) + thus ?thesis + using assms card_of_Field_ordIso[of r] + ordIso_transitive[of "|Field r <+> Field p|"] + ordIso_transitive[of _ "|Field r|"] by blast +qed + + +subsection {* The cardinal $\omega$ and the finite cardinals *} + + +text{* The cardinal $\omega$, of natural numbers, shall be the standard non-strict +order relation on +@{text "nat"}, that we abbreviate by @{text "natLeq"}. The finite cardinals +shall be the restrictions of these relations to the numbers smaller than +fixed numbers @{text "n"}, that we abbreviate by @{text "natLeq_on n"}. *} + +abbreviation "(natLeq::(nat * nat) set) \ {(x,y). x \ y}" +abbreviation "(natLess::(nat * nat) set) \ {(x,y). x < y}" + +abbreviation natLeq_on :: "nat \ (nat * nat) set" +where "natLeq_on n \ {(x,y). x < n \ y < n \ x \ y}" + +lemma infinite_cartesian_product: +assumes "infinite A" "infinite B" +shows "infinite (A \ B)" +proof + assume "finite (A \ B)" + from assms(1) have "A \ {}" by auto + with `finite (A \ B)` have "finite B" using finite_cartesian_productD2 by auto + with assms(2) show False by simp +qed + + +subsubsection {* First as well-orders *} + + +lemma Field_natLeq: "Field natLeq = (UNIV::nat set)" +by(unfold Field_def, auto) + + +lemma natLeq_Refl: "Refl natLeq" +unfolding refl_on_def Field_def by auto + + +lemma natLeq_trans: "trans natLeq" +unfolding trans_def by auto + + +lemma natLeq_Preorder: "Preorder natLeq" +unfolding preorder_on_def +by (auto simp add: natLeq_Refl natLeq_trans) + + +lemma natLeq_antisym: "antisym natLeq" +unfolding antisym_def by auto + + +lemma natLeq_Partial_order: "Partial_order natLeq" +unfolding partial_order_on_def +by (auto simp add: natLeq_Preorder natLeq_antisym) + + +lemma natLeq_Total: "Total natLeq" +unfolding total_on_def by auto + + +lemma natLeq_Linear_order: "Linear_order natLeq" +unfolding linear_order_on_def +by (auto simp add: natLeq_Partial_order natLeq_Total) + + +lemma natLeq_natLess_Id: "natLess = natLeq - Id" +by auto + + +lemma natLeq_Well_order: "Well_order natLeq" +unfolding well_order_on_def +using natLeq_Linear_order wf_less natLeq_natLess_Id by auto + + +lemma closed_nat_set_iff: +assumes "\(m::nat) n. n \ A \ m \ n \ m \ A" +shows "A = UNIV \ (\n. A = {0 ..< n})" +proof- + {assume "A \ UNIV" hence "\n. n \ A" by blast + moreover obtain n where n_def: "n = (LEAST n. n \ A)" by blast + ultimately have 1: "n \ A \ (\m. m < n \ m \ A)" + using LeastI_ex[of "\ n. n \ A"] n_def Least_le[of "\ n. n \ A"] by fastforce + have "A = {0 ..< n}" + proof(auto simp add: 1) + fix m assume *: "m \ A" + {assume "n \ m" with assms * have "n \ A" by blast + hence False using 1 by auto + } + thus "m < n" by fastforce + qed + hence "\n. A = {0 ..< n}" by blast + } + thus ?thesis by blast +qed + + +lemma Field_natLeq_on: "Field (natLeq_on n) = {0 ..< n}" +unfolding Field_def by auto + + +lemma natLeq_underS_less: "rel.underS natLeq n = {0 ..< n}" +unfolding rel.underS_def by auto + + +lemma Restr_natLeq: "Restr natLeq {0 ..< n} = natLeq_on n" +by auto + + +lemma Restr_natLeq2: +"Restr natLeq (rel.underS natLeq n) = natLeq_on n" +by (auto simp add: Restr_natLeq natLeq_underS_less) + + +lemma natLeq_on_Well_order: "Well_order(natLeq_on n)" +using Restr_natLeq[of n] natLeq_Well_order + Well_order_Restr[of natLeq "{0.. m \ wo_rel.ofilter (natLeq_on m) {0 ..< n}" +by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def, + simp add: Field_natLeq_on, unfold rel.under_def, auto) + + +lemma natLeq_on_ofilter_iff: +"wo_rel.ofilter (natLeq_on m) A = (\n \ m. A = {0 ..< n})" +proof(rule iffI) + assume *: "wo_rel.ofilter (natLeq_on m) A" + hence 1: "A \ {0..n1 n2. n2 \ A \ n1 \ n2 \ n1 \ A" + using * by(fastforce simp add: natLeq_on_wo_rel wo_rel.ofilter_def rel.under_def) + hence "A = UNIV \ (\n. A = {0 ..< n})" using closed_nat_set_iff by blast + thus "\n \ m. A = {0 ..< n}" using 1 atLeastLessThan_less_eq by blast +next + assume "(\n\m. A = {0 ..< n})" + thus "wo_rel.ofilter (natLeq_on m) A" by (auto simp add: natLeq_on_ofilter_less_eq) +qed + + + +subsubsection {* Then as cardinals *} + + +lemma natLeq_Card_order: "Card_order natLeq" +proof(auto simp add: natLeq_Well_order + Card_order_iff_Restr_underS Restr_natLeq2, simp add: Field_natLeq) + fix n have "finite(Field (natLeq_on n))" + unfolding Field_natLeq_on by auto + moreover have "infinite(UNIV::nat set)" by auto + ultimately show "natLeq_on n o |A| )" +using infinite_iff_card_of_nat[of A] card_of_nat + ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast + + +corollary finite_iff_ordLess_natLeq: +"finite A = ( |A| finite A" +unfolding ordIso_def iso_def[abs_def] +by (auto simp: Field_natLeq_on bij_betw_finite Field_card_of) + + +lemma natLeq_on_Card_order: "Card_order (natLeq_on n)" +proof(unfold card_order_on_def, + auto simp add: natLeq_on_Well_order, simp add: Field_natLeq_on) + fix r assume "well_order_on {0..o r" + using finite_atLeastLessThan natLeq_on_well_order_on + finite_well_order_on_ordIso ordIso_iff_ordLeq by blast +qed + + +corollary card_of_Field_natLeq_on: +"|Field (natLeq_on n)| =o natLeq_on n" +using Field_natLeq_on natLeq_on_Card_order + Card_order_iff_ordIso_card_of[of "natLeq_on n"] + ordIso_symmetric[of "natLeq_on n"] by blast + + +corollary card_of_less: +"|{0 ..< n}| =o natLeq_on n" +using Field_natLeq_on card_of_Field_natLeq_on by auto + + +lemma natLeq_on_ordLeq_less_eq: +"((natLeq_on m) \o (natLeq_on n)) = (m \ n)" +proof + assume "natLeq_on m \o natLeq_on n" + then obtain f where "inj_on f {0.. f ` {0.. {0.. n" using atLeastLessThan_less_eq2 by blast +next + assume "m \ n" + hence "inj_on id {0.. id ` {0.. {0..o |{0..o natLeq_on n" + using card_of_less ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast +qed + + +lemma natLeq_on_ordLeq_less: +"((natLeq_on m) o |B| ) = (card A \ card B)" +using assms card_of_ordLeq[of A B] inj_on_iff_card_le[of A B] by blast + + +lemma finite_imp_card_of_natLeq_on: +assumes "finite A" +shows "|A| =o natLeq_on (card A)" +proof- + obtain h where "bij_betw h A {0 ..< card A}" + using assms ex_bij_betw_finite_nat by blast + thus ?thesis using card_of_ordIso card_of_less ordIso_equivalence by blast +qed + + +lemma finite_iff_card_of_natLeq_on: +"finite A = (\n. |A| =o natLeq_on n)" +using finite_imp_card_of_natLeq_on[of A] +by(auto simp add: ordIso_natLeq_on_imp_finite) + + + +subsection {* The successor of a cardinal *} + + +text{* First we define @{text "isCardSuc r r'"}, the notion of @{text "r'"} +being a successor cardinal of @{text "r"}. Although the definition does +not require @{text "r"} to be a cardinal, only this case will be meaningful. *} + + +definition isCardSuc :: "'a rel \ 'a set rel \ bool" +where +"isCardSuc r r' \ + Card_order r' \ r + (\(r''::'a set rel). Card_order r'' \ r r' \o r'')" + + +text{* Now we introduce the cardinal-successor operator @{text "cardSuc"}, +by picking {\em some} cardinal-order relation fulfilling @{text "isCardSuc"}. +Again, the picked item shall be proved unique up to order-isomorphism. *} + + +definition cardSuc :: "'a rel \ 'a set rel" +where +"cardSuc r \ SOME r'. isCardSuc r r'" + + +lemma exists_minim_Card_order: +"\R \ {}; \r \ R. Card_order r\ \ \r \ R. \r' \ R. r \o r'" +unfolding card_order_on_def using exists_minim_Well_order by blast + + +lemma exists_isCardSuc: +assumes "Card_order r" +shows "\r'. isCardSuc r r'" +proof- + let ?R = "{(r'::'a set rel). Card_order r' \ r ?R \ (\r \ ?R. Card_order r)" using assms + by (simp add: card_of_Card_order Card_order_Pow) + then obtain r where "r \ ?R \ (\r' \ ?R. r \o r')" + using exists_minim_Card_order[of ?R] by blast + thus ?thesis unfolding isCardSuc_def by auto +qed + + +lemma cardSuc_isCardSuc: +assumes "Card_order r" +shows "isCardSuc r (cardSuc r)" +unfolding cardSuc_def using assms +by (simp add: exists_isCardSuc someI_ex) + + +lemma cardSuc_Card_order: +"Card_order r \ Card_order(cardSuc r)" +using cardSuc_isCardSuc unfolding isCardSuc_def by blast + + +lemma cardSuc_greater: +"Card_order r \ r r \o cardSuc r" +using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast + + +text{* The minimality property of @{text "cardSuc"} originally present in its definition +is local to the type @{text "'a set rel"}, i.e., that of @{text "cardSuc r"}: *} + +lemma cardSuc_least_aux: +"\Card_order (r::'a rel); Card_order (r'::'a set rel); r \ cardSuc r \o r'" +using cardSuc_isCardSuc unfolding isCardSuc_def by blast + + +text{* But from this we can infer general minimality: *} + +lemma cardSuc_least: +assumes CARD: "Card_order r" and CARD': "Card_order r'" and LESS: "r o r'" +proof- + let ?p = "cardSuc r" + have 0: "Well_order ?p \ Well_order r'" + using assms cardSuc_Card_order unfolding card_order_on_def by blast + {assume "r' r'' o r''" using cardSuc_least_aux CARD by blast + hence False using 2 not_ordLess_ordLeq by blast + } + thus ?thesis using 0 ordLess_or_ordLeq by blast +qed + + +lemma cardSuc_ordLess_ordLeq: +assumes CARD: "Card_order r" and CARD': "Card_order r'" +shows "(r o r')" +proof(auto simp add: assms cardSuc_least) + assume "cardSuc r \o r'" + thus "r o r)" +proof- + have "Well_order r \ Well_order r'" + using assms unfolding card_order_on_def by auto + moreover have "Well_order(cardSuc r)" + using assms cardSuc_Card_order card_order_on_def by blast + ultimately show ?thesis + using assms cardSuc_ordLess_ordLeq[of r r'] + not_ordLeq_iff_ordLess[of r r'] not_ordLeq_iff_ordLess[of r' "cardSuc r"] by blast +qed + + +lemma cardSuc_mono_ordLeq: +assumes CARD: "Card_order r" and CARD': "Card_order r'" +shows "(cardSuc r \o cardSuc r') = (r \o r')" +using assms cardSuc_ordLeq_ordLess cardSuc_ordLess_ordLeq cardSuc_Card_order by blast + + +lemma cardSuc_invar_ordIso: +assumes CARD: "Card_order r" and CARD': "Card_order r'" +shows "(cardSuc r =o cardSuc r') = (r =o r')" +proof- + have 0: "Well_order r \ Well_order r' \ Well_order(cardSuc r) \ Well_order(cardSuc r')" + using assms by (simp add: card_order_on_well_order_on cardSuc_Card_order) + thus ?thesis + using ordIso_iff_ordLeq[of r r'] ordIso_iff_ordLeq + using cardSuc_mono_ordLeq[of r r'] cardSuc_mono_ordLeq[of r' r] assms by blast +qed + + +lemma cardSuc_natLeq_on_Suc: +"cardSuc(natLeq_on n) =o natLeq_on(Suc n)" +proof- + obtain r r' p where r_def: "r = natLeq_on n" and + r'_def: "r' = cardSuc(natLeq_on n)" and + p_def: "p = natLeq_on(Suc n)" by blast + (* Preliminary facts: *) + have CARD: "Card_order r \ Card_order r' \ Card_order p" unfolding r_def r'_def p_def + using cardSuc_ordLess_ordLeq natLeq_on_Card_order cardSuc_Card_order by blast + hence WELL: "Well_order r \ Well_order r' \ Well_order p" + unfolding card_order_on_def by force + have FIELD: "Field r = {0.. Field p = {0..<(Suc n)}" + unfolding r_def p_def Field_natLeq_on by simp + hence FIN: "finite (Field r)" by force + have "r o p" using CARD unfolding r_def r'_def p_def + using natLeq_on_ordLeq_less cardSuc_ordLess_ordLeq by blast + moreover have "p \o r'" + proof- + {assume "r' Suc n" and 3: "f ` (Field r') = {0.. n" using 2 by force + (* *) + have "bij_betw f (Field r') (f ` (Field r'))" + using 1 WELL embed_inj_on unfolding bij_betw_def by force + moreover have "finite(f ` (Field r'))" using 3 finite_atLeastLessThan[of 0 m] by force + ultimately have 5: "finite (Field r') \ card(Field r') = card (f ` (Field r'))" + using bij_betw_same_card bij_betw_finite by metis + hence "card(Field r') \ card(Field r)" using 3 4 FIELD by force + hence "|Field r'| \o |Field r|" using FIN 5 finite_card_of_iff_card2 by blast + hence False using LESS not_ordLess_ordLeq by auto + } + thus ?thesis using WELL CARD by (fastforce simp: not_ordLess_iff_ordLeq) + qed + ultimately show ?thesis using ordIso_iff_ordLeq unfolding r'_def p_def by blast +qed + + +lemma card_of_cardSuc_finite: +"finite(Field(cardSuc |A| )) = finite A" +proof + assume *: "finite (Field (cardSuc |A| ))" + have 0: "|Field(cardSuc |A| )| =o cardSuc |A|" + using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast + hence "|A| \o |Field(cardSuc |A| )|" + using card_of_Card_order[of A] cardSuc_ordLeq[of "|A|"] ordIso_symmetric + ordLeq_ordIso_trans by blast + thus "finite A" using * card_of_ordLeq_finite by blast +next + assume "finite A" + then obtain n where "|A| =o natLeq_on n" using finite_iff_card_of_natLeq_on by blast + hence "cardSuc |A| =o cardSuc(natLeq_on n)" + using card_of_Card_order cardSuc_invar_ordIso natLeq_on_Card_order by blast + hence "cardSuc |A| =o natLeq_on(Suc n)" + using cardSuc_natLeq_on_Suc ordIso_transitive by blast + hence "cardSuc |A| =o |{0..<(Suc n)}|" using card_of_less ordIso_equivalence by blast + moreover have "|Field (cardSuc |A| ) | =o cardSuc |A|" + using card_of_Field_ordIso cardSuc_Card_order card_of_Card_order by blast + ultimately have "|Field (cardSuc |A| ) | =o |{0..<(Suc n)}|" + using ordIso_equivalence by blast + thus "finite (Field (cardSuc |A| ))" + using card_of_ordIso_finite finite_atLeastLessThan by blast +qed + + +lemma cardSuc_finite: +assumes "Card_order r" +shows "finite (Field (cardSuc r)) = finite (Field r)" +proof- + let ?A = "Field r" + have "|?A| =o r" using assms by (simp add: card_of_Field_ordIso) + hence "cardSuc |?A| =o cardSuc r" using assms + by (simp add: card_of_Card_order cardSuc_invar_ordIso) + moreover have "|Field (cardSuc |?A| ) | =o cardSuc |?A|" + by (simp add: card_of_card_order_on Field_card_of card_of_Field_ordIso cardSuc_Card_order) + moreover + {have "|Field (cardSuc r) | =o cardSuc r" + using assms by (simp add: card_of_Field_ordIso cardSuc_Card_order) + hence "cardSuc r =o |Field (cardSuc r) |" + using ordIso_symmetric by blast + } + ultimately have "|Field (cardSuc |?A| ) | =o |Field (cardSuc r) |" + using ordIso_transitive by blast + hence "finite (Field (cardSuc |?A| )) = finite (Field (cardSuc r))" + using card_of_ordIso_finite by blast + thus ?thesis by (simp only: card_of_cardSuc_finite) +qed + + +lemma card_of_Plus_ordLess_infinite: +assumes INF: "infinite C" and + LESS1: "|A| B| B = {}") + assume Case1: "A = {} \ B = {}" + hence "|A| =o |A <+> B| \ |B| =o |A <+> B|" + using card_of_Plus_empty1 card_of_Plus_empty2 by blast + hence "|A <+> B| =o |A| \ |A <+> B| =o |B|" + using ordIso_symmetric[of "|A|"] ordIso_symmetric[of "|B|"] by blast + thus ?thesis using LESS1 LESS2 + ordIso_ordLess_trans[of "|A <+> B|" "|A|"] + ordIso_ordLess_trans[of "|A <+> B|" "|B|"] by blast +next + assume Case2: "\(A = {} \ B = {})" + {assume *: "|C| \o |A <+> B|" + hence "infinite (A <+> B)" using INF card_of_ordLeq_finite by blast + hence 1: "infinite A \ infinite B" using finite_Plus by blast + {assume Case21: "|A| \o |B|" + hence "infinite B" using 1 card_of_ordLeq_finite by blast + hence "|A <+> B| =o |B|" using Case2 Case21 + by (auto simp add: card_of_Plus_infinite) + hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast + } + moreover + {assume Case22: "|B| \o |A|" + hence "infinite A" using 1 card_of_ordLeq_finite by blast + hence "|A <+> B| =o |A|" using Case2 Case22 + by (auto simp add: card_of_Plus_infinite) + hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast + } + ultimately have False using ordLeq_total card_of_Well_order[of A] + card_of_Well_order[of B] by blast + } + thus ?thesis using ordLess_or_ordLeq[of "|A <+> B|" "|C|"] + card_of_Well_order[of "A <+> B"] card_of_Well_order[of "C"] by auto +qed + + +lemma card_of_Plus_ordLess_infinite_Field: +assumes INF: "infinite (Field r)" and r: "Card_order r" and + LESS1: "|A| B| |?C| =o r" using r card_of_Field_ordIso + ordIso_symmetric by blast + hence "|A| B| o r" and B: "|B| \o r" +and c: "Card_order r" +shows "|A <+> B| \o r" +proof- + let ?r' = "cardSuc r" + have "Card_order ?r' \ infinite (Field ?r')" using assms + by (simp add: cardSuc_Card_order cardSuc_finite) + moreover have "|A| B| o r" and B: "|B| \o r" +and "Card_order r" +shows "|A Un B| \o r" +using assms card_of_Plus_ordLeq_infinite_Field card_of_Un_Plus_ordLeq +ordLeq_transitive by blast + + + +subsection {* Regular cardinals *} + + +definition cofinal where +"cofinal A r \ + ALL a : Field r. EX b : A. a \ b \ (a,b) : r" + + +definition regular where +"regular r \ + ALL K. K \ Field r \ cofinal K r \ |K| =o r" + + +definition relChain where +"relChain r As \ + ALL i j. (i,j) \ r \ As i \ As j" + +lemma regular_UNION: +assumes r: "Card_order r" "regular r" +and As: "relChain r As" +and Bsub: "B \ (UN i : Field r. As i)" +and cardB: "|B| As i" +proof- + let ?phi = "%b j. j : Field r \ b : As j" + have "ALL b : B. EX j. ?phi b j" using Bsub by blast + then obtain f where f: "!! b. b : B \ ?phi b (f b)" + using bchoice[of B ?phi] by blast + let ?K = "f ` B" + {assume 1: "!! i. i : Field r \ ~ B \ As i" + have 2: "cofinal ?K r" + unfolding cofinal_def proof auto + fix i assume i: "i : Field r" + with 1 obtain b where b: "b : B \ b \ As i" by blast + hence "i \ f b \ ~ (f b,i) : r" + using As f unfolding relChain_def by auto + hence "i \ f b \ (i, f b) : r" using r + unfolding card_order_on_def well_order_on_def linear_order_on_def + total_on_def using i f b by auto + with b show "\b\B. i \ f b \ (i, f b) \ r" by blast + qed + moreover have "?K \ Field r" using f by blast + ultimately have "|?K| =o r" using 2 r unfolding regular_def by blast + moreover + { + have "|?K| <=o |B|" using card_of_image . + hence "|?K| (p \o r) = (p Field ?r'" and 2: "cofinal K ?r'" + hence "|K| \o |Field ?r'|" by (simp only: card_of_mono1) + also have 22: "|Field ?r'| =o ?r'" + using r' by (simp add: card_of_Field_ordIso[of ?r']) + finally have "|K| \o ?r'" . + moreover + {let ?L = "UN j : K. rel.underS ?r' j" + let ?J = "Field r" + have rJ: "r =o |?J|" + using r_card card_of_Field_ordIso ordIso_symmetric by blast + assume "|K| o |?J|" using rJ ordLeq_ordIso_trans by blast + moreover + {have "ALL j : K. |rel.underS ?r' j| o r" + using r' card_of_Card_order by blast + hence "ALL j : K. |rel.underS ?r' j| \o |?J|" + using rJ ordLeq_ordIso_trans by blast + } + ultimately have "|?L| \o |?J|" + using r_inf card_of_UNION_ordLeq_infinite by blast + hence "|?L| \o r" using rJ ordIso_symmetric ordLeq_ordIso_trans by blast + hence "|?L| ?L" + using 2 unfolding rel.underS_def cofinal_def by auto + hence "|Field ?r'| \o |?L|" by (simp add: card_of_mono1) + hence "?r' \o |?L|" + using 22 ordIso_ordLeq_trans ordIso_symmetric by blast + } + ultimately have "|?L| (UN i : Field (cardSuc r). As i)" +and cardB: "|B| <=o r" +shows "EX i : Field (cardSuc r). B \ As i" +proof- + let ?r' = "cardSuc r" + have "Card_order ?r' \ |B| a \ A. f a \ B) \ (\ a. a \ A \ f a = undefined)}" + +lemma Func_empty: +"Func {} B = {\x. undefined}" +unfolding Func_def by auto + +lemma Func_elim: +assumes "g \ Func A B" and "a \ A" +shows "\ b. b \ B \ g a = b" +using assms unfolding Func_def by (cases "g a = undefined") auto + +definition curr where +"curr A f \ \ a. if a \ A then \b. f (a,b) else undefined" + +lemma curr_in: +assumes f: "f \ Func (A <*> B) C" +shows "curr A f \ Func A (Func B C)" +using assms unfolding curr_def Func_def by auto + +lemma curr_inj: +assumes "f1 \ Func (A <*> B) C" and "f2 \ Func (A <*> B) C" +shows "curr A f1 = curr A f2 \ f1 = f2" +proof safe + assume c: "curr A f1 = curr A f2" + show "f1 = f2" + proof (rule ext, clarify) + fix a b show "f1 (a, b) = f2 (a, b)" + proof (cases "(a,b) \ A <*> B") + case False + thus ?thesis using assms unfolding Func_def by auto + next + case True hence a: "a \ A" and b: "b \ B" by auto + thus ?thesis + using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp + qed + qed +qed + +lemma curr_surj: +assumes "g \ Func A (Func B C)" +shows "\ f \ Func (A <*> B) C. curr A f = g" +proof + let ?f = "\ ab. if fst ab \ A \ snd ab \ B then g (fst ab) (snd ab) else undefined" + show "curr A ?f = g" + proof (rule ext) + fix a show "curr A ?f a = g a" + proof (cases "a \ A") + case False + hence "g a = undefined" using assms unfolding Func_def by auto + thus ?thesis unfolding curr_def using False by simp + next + case True + obtain g1 where "g1 \ Func B C" and "g a = g1" + using assms using Func_elim[OF assms True] by blast + thus ?thesis using True unfolding Func_def curr_def by auto + qed + qed + show "?f \ Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto +qed + +lemma bij_betw_curr: +"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))" +unfolding bij_betw_def inj_on_def image_def +using curr_in curr_inj curr_surj by blast + +lemma card_of_Func_Times: +"|Func (A <*> B) C| =o |Func A (Func B C)|" +unfolding card_of_ordIso[symmetric] +using bij_betw_curr by blast + +definition Func_map where +"Func_map B2 f1 f2 g b2 \ if b2 \ B2 then f1 (g (f2 b2)) else undefined" + +lemma Func_map: +assumes g: "g \ Func A2 A1" and f1: "f1 ` A1 \ B1" and f2: "f2 ` B2 \ A2" +shows "Func_map B2 f1 f2 g \ Func B2 B1" +using assms unfolding Func_def Func_map_def mem_Collect_eq by auto + +lemma Func_non_emp: +assumes "B \ {}" +shows "Func A B \ {}" +proof- + obtain b where b: "b \ B" using assms by auto + hence "(\ a. if a \ A then b else undefined) \ Func A B" unfolding Func_def by auto + thus ?thesis by blast +qed + +lemma Func_is_emp: +"Func A B = {} \ A \ {} \ B = {}" (is "?L \ ?R") +proof + assume L: ?L + moreover {assume "A = {}" hence False using L Func_empty by auto} + moreover {assume "B \ {}" hence False using L Func_non_emp by metis} + ultimately show ?R by blast +next + assume R: ?R + moreover + {fix f assume "f \ Func A B" + moreover obtain a where "a \ A" using R by blast + ultimately obtain b where "b \ B" unfolding Func_def by(cases "f a = undefined", force+) + with R have False by auto + } + thus ?L by blast +qed + +lemma Func_map_surj: +assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \ A2" +and B2A2: "B2 = {} \ A2 = {}" +shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1" +proof(cases "B2 = {}") + case True + thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def) +next + case False note B2 = False + show ?thesis + proof safe + fix h assume h: "h \ Func B2 B1" + def j1 \ "inv_into A1 f1" + have "\ a2 \ f2 ` B2. \ b2. b2 \ B2 \ f2 b2 = a2" by blast + then obtain k where k: "\ a2 \ f2 ` B2. k a2 \ B2 \ f2 (k a2) = a2" by metis + {fix b2 assume b2: "b2 \ B2" + hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto + moreover have "k (f2 b2) \ B2" using b2 A2(2) k by auto + ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast + } note kk = this + obtain b22 where b22: "b22 \ B2" using B2 by auto + def j2 \ "\ a2. if a2 \ f2 ` B2 then k a2 else b22" + have j2A2: "j2 ` A2 \ B2" unfolding j2_def using k b22 by auto + have j2: "\ b2. b2 \ B2 \ j2 (f2 b2) = b2" + using kk unfolding j2_def by auto + def g \ "Func_map A2 j1 j2 h" + have "Func_map B2 f1 f2 g = h" + proof (rule ext) + fix b2 show "Func_map B2 f1 f2 g b2 = h b2" + proof(cases "b2 \ B2") + case True + show ?thesis + proof (cases "h b2 = undefined") + case True + hence b1: "h b2 \ f1 ` A1" using h `b2 \ B2` unfolding B1 Func_def by auto + show ?thesis using A2 f_inv_into_f[OF b1] + unfolding True g_def Func_map_def j1_def j2[OF `b2 \ B2`] by auto + qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def, + auto intro: f_inv_into_f) + qed(insert h, unfold Func_def Func_map_def, auto) + qed + moreover have "g \ Func A2 A1" unfolding g_def apply(rule Func_map[OF h]) + using inv_into_into j2A2 B1 A2 inv_into_into + unfolding j1_def image_def by fast+ + ultimately show "h \ Func_map B2 f1 f2 ` Func A2 A1" + unfolding Func_map_def[abs_def] unfolding image_def by auto + qed(insert B1 Func_map[OF _ _ A2(2)], auto) +qed + +lemma card_of_Pow_Func: +"|Pow A| =o |Func A (UNIV::bool set)|" +proof- + def F \ "\ A' a. if a \ A then (if a \ A' then True else False) + else undefined" + have "bij_betw F (Pow A) (Func A (UNIV::bool set))" + unfolding bij_betw_def inj_on_def proof (intro ballI impI conjI) + fix A1 A2 assume "A1 \ Pow A" "A2 \ Pow A" "F A1 = F A2" + thus "A1 = A2" unfolding F_def Pow_def fun_eq_iff by (auto split: split_if_asm) + next + show "F ` Pow A = Func A UNIV" + proof safe + fix f assume f: "f \ Func A (UNIV::bool set)" + show "f \ F ` Pow A" unfolding image_def mem_Collect_eq proof(intro bexI) + let ?A1 = "{a \ A. f a = True}" + show "f = F ?A1" unfolding F_def apply(rule ext) + using f unfolding Func_def mem_Collect_eq by auto + qed auto + qed(unfold Func_def mem_Collect_eq F_def, auto) + qed + thus ?thesis unfolding card_of_ordIso[symmetric] by blast +qed + +lemma card_of_Func_UNIV: +"|Func (UNIV::'a set) (B::'b set)| =o |{f::'a \ 'b. range f \ B}|" +apply(rule ordIso_symmetric) proof(intro card_of_ordIsoI) + let ?F = "\ f (a::'a). ((f a)::'b)" + show "bij_betw ?F {f. range f \ B} (Func UNIV B)" + unfolding bij_betw_def inj_on_def proof safe + fix h :: "'a \ 'b" assume h: "h \ Func UNIV B" + hence "\ a. \ b. h a = b" unfolding Func_def by auto + then obtain f where f: "\ a. h a = f a" by metis + hence "range f \ B" using h unfolding Func_def by auto + thus "h \ (\f a. f a) ` {f. range f \ B}" using f unfolding image_def by auto + qed(unfold Func_def fun_eq_iff, auto) +qed + +end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Cardinal_Order_Relation_LFP.thy --- a/src/HOL/Cardinals/Cardinal_Order_Relation_LFP.thy Mon Nov 18 18:04:45 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2161 +0,0 @@ -(* Title: HOL/Cardinals/Cardinal_Order_Relation_LFP.thy - Author: Andrei Popescu, TU Muenchen - Copyright 2012 - -Cardinal-order relations (LFP). -*) - -header {* Cardinal-Order Relations (LFP) *} - -theory Cardinal_Order_Relation_LFP -imports Constructions_on_Wellorders_LFP -begin - - -text{* In this section, we define cardinal-order relations to be minim well-orders -on their field. Then we define the cardinal of a set to be {\em some} cardinal-order -relation on that set, which will be unique up to order isomorphism. Then we study -the connection between cardinals and: -\begin{itemize} -\item standard set-theoretic constructions: products, -sums, unions, lists, powersets, set-of finite sets operator; -\item finiteness and infiniteness (in particular, with the numeric cardinal operator -for finite sets, @{text "card"}, from the theory @{text "Finite_Sets.thy"}). -\end{itemize} -% -On the way, we define the canonical $\omega$ cardinal and finite cardinals. We also -define (again, up to order isomorphism) the successor of a cardinal, and show that -any cardinal admits a successor. - -Main results of this section are the existence of cardinal relations and the -facts that, in the presence of infiniteness, -most of the standard set-theoretic constructions (except for the powerset) -{\em do not increase cardinality}. In particular, e.g., the set of words/lists over -any infinite set has the same cardinality (hence, is in bijection) with that set. -*} - - -subsection {* Cardinal orders *} - - -text{* A cardinal order in our setting shall be a well-order {\em minim} w.r.t. the -order-embedding relation, @{text "\o"} (which is the same as being {\em minimal} w.r.t. the -strict order-embedding relation, @{text " 'a rel \ bool" -where -"card_order_on A r \ well_order_on A r \ (\r'. well_order_on A r' \ r \o r')" - - -abbreviation "Card_order r \ card_order_on (Field r) r" -abbreviation "card_order r \ card_order_on UNIV r" - - -lemma card_order_on_well_order_on: -assumes "card_order_on A r" -shows "well_order_on A r" -using assms unfolding card_order_on_def by simp - - -lemma card_order_on_Card_order: -"card_order_on A r \ A = Field r \ Card_order r" -unfolding card_order_on_def using rel.well_order_on_Field by blast - - -text{* The existence of a cardinal relation on any given set (which will mean -that any set has a cardinal) follows from two facts: -\begin{itemize} -\item Zermelo's theorem (proved in @{text "Zorn.thy"} as theorem @{text "well_order_on"}), -which states that on any given set there exists a well-order; -\item The well-founded-ness of @{text "r. card_order_on A r" -proof- - obtain R where R_def: "R = {r. well_order_on A r}" by blast - have 1: "R \ {} \ (\r \ R. Well_order r)" - using well_order_on[of A] R_def rel.well_order_on_Well_order by blast - hence "\r \ R. \r' \ R. r \o r'" - using exists_minim_Well_order[of R] by auto - thus ?thesis using R_def unfolding card_order_on_def by auto -qed - - -lemma card_order_on_ordIso: -assumes CO: "card_order_on A r" and CO': "card_order_on A r'" -shows "r =o r'" -using assms unfolding card_order_on_def -using ordIso_iff_ordLeq by blast - - -lemma Card_order_ordIso: -assumes CO: "Card_order r" and ISO: "r' =o r" -shows "Card_order r'" -using ISO unfolding ordIso_def -proof(unfold card_order_on_def, auto) - fix p' assume "well_order_on (Field r') p'" - hence 0: "Well_order p' \ Field p' = Field r'" - using rel.well_order_on_Well_order by blast - obtain f where 1: "iso r' r f" and 2: "Well_order r \ Well_order r'" - using ISO unfolding ordIso_def by auto - hence 3: "inj_on f (Field r') \ f ` (Field r') = Field r" - by (auto simp add: iso_iff embed_inj_on) - let ?p = "dir_image p' f" - have 4: "p' =o ?p \ Well_order ?p" - using 0 2 3 by (auto simp add: dir_image_ordIso Well_order_dir_image) - moreover have "Field ?p = Field r" - using 0 3 by (auto simp add: dir_image_Field2 order_on_defs) - ultimately have "well_order_on (Field r) ?p" by auto - hence "r \o ?p" using CO unfolding card_order_on_def by auto - thus "r' \o p'" - using ISO 4 ordLeq_ordIso_trans ordIso_ordLeq_trans ordIso_symmetric by blast -qed - - -lemma Card_order_ordIso2: -assumes CO: "Card_order r" and ISO: "r =o r'" -shows "Card_order r'" -using assms Card_order_ordIso ordIso_symmetric by blast - - -subsection {* Cardinal of a set *} - - -text{* We define the cardinal of set to be {\em some} cardinal order on that set. -We shall prove that this notion is unique up to order isomorphism, meaning -that order isomorphism shall be the true identity of cardinals. *} - - -definition card_of :: "'a set \ 'a rel" ("|_|" ) -where "card_of A = (SOME r. card_order_on A r)" - - -lemma card_of_card_order_on: "card_order_on A |A|" -unfolding card_of_def by (auto simp add: card_order_on someI_ex) - - -lemma card_of_well_order_on: "well_order_on A |A|" -using card_of_card_order_on card_order_on_def by blast - - -lemma Field_card_of: "Field |A| = A" -using card_of_card_order_on[of A] unfolding card_order_on_def -using rel.well_order_on_Field by blast - - -lemma card_of_Card_order: "Card_order |A|" -by (simp only: card_of_card_order_on Field_card_of) - - -corollary ordIso_card_of_imp_Card_order: -"r =o |A| \ Card_order r" -using card_of_Card_order Card_order_ordIso by blast - - -lemma card_of_Well_order: "Well_order |A|" -using card_of_Card_order unfolding card_order_on_def by auto - - -lemma card_of_refl: "|A| =o |A|" -using card_of_Well_order ordIso_reflexive by blast - - -lemma card_of_least: "well_order_on A r \ |A| \o r" -using card_of_card_order_on unfolding card_order_on_def by blast - - -lemma card_of_ordIso: -"(\f. bij_betw f A B) = ( |A| =o |B| )" -proof(auto) - fix f assume *: "bij_betw f A B" - then obtain r where "well_order_on B r \ |A| =o r" - using Well_order_iso_copy card_of_well_order_on by blast - hence "|B| \o |A|" using card_of_least - ordLeq_ordIso_trans ordIso_symmetric by blast - moreover - {let ?g = "inv_into A f" - have "bij_betw ?g B A" using * bij_betw_inv_into by blast - then obtain r where "well_order_on A r \ |B| =o r" - using Well_order_iso_copy card_of_well_order_on by blast - hence "|A| \o |B|" using card_of_least - ordLeq_ordIso_trans ordIso_symmetric by blast - } - ultimately show "|A| =o |B|" using ordIso_iff_ordLeq by blast -next - assume "|A| =o |B|" - then obtain f where "iso ( |A| ) ( |B| ) f" - unfolding ordIso_def by auto - hence "bij_betw f A B" unfolding iso_def Field_card_of by simp - thus "\f. bij_betw f A B" by auto -qed - - -lemma card_of_ordLeq: -"(\f. inj_on f A \ f ` A \ B) = ( |A| \o |B| )" -proof(auto) - fix f assume *: "inj_on f A" and **: "f ` A \ B" - {assume "|B| o |A|" using ordLeq_iff_ordLess_or_ordIso by blast - then obtain g where "embed ( |B| ) ( |A| ) g" - unfolding ordLeq_def by auto - hence 1: "inj_on g B \ g ` B \ A" using embed_inj_on[of "|B|" "|A|" "g"] - card_of_Well_order[of "B"] Field_card_of[of "B"] Field_card_of[of "A"] - embed_Field[of "|B|" "|A|" g] by auto - obtain h where "bij_betw h A B" - using * ** 1 Cantor_Bernstein[of f] by fastforce - hence "|A| =o |B|" using card_of_ordIso by blast - hence "|A| \o |B|" using ordIso_iff_ordLeq by auto - } - thus "|A| \o |B|" using ordLess_or_ordLeq[of "|B|" "|A|"] - by (auto simp: card_of_Well_order) -next - assume *: "|A| \o |B|" - obtain f where "embed ( |A| ) ( |B| ) f" - using * unfolding ordLeq_def by auto - hence "inj_on f A \ f ` A \ B" using embed_inj_on[of "|A|" "|B|" f] - card_of_Well_order[of "A"] Field_card_of[of "A"] Field_card_of[of "B"] - embed_Field[of "|A|" "|B|" f] by auto - thus "\f. inj_on f A \ f ` A \ B" by auto -qed - - -lemma card_of_ordLeq2: -"A \ {} \ (\g. g ` B = A) = ( |A| \o |B| )" -using card_of_ordLeq[of A B] inj_on_iff_surj[of A B] by auto - - -lemma card_of_ordLess: -"(\(\f. inj_on f A \ f ` A \ B)) = ( |B| (\f. inj_on f A \ f ` A \ B)) = (\ |A| \o |B| )" - using card_of_ordLeq by blast - also have "\ = ( |B| {} \ (\(\f. f ` A = B)) = ( |A| a. a \ A \ f a \ B" -shows "|A| \o |B|" -using assms unfolding card_of_ordLeq[symmetric] by auto - - -lemma card_of_unique: -"card_order_on A r \ r =o |A|" -by (simp only: card_order_on_ordIso card_of_card_order_on) - - -lemma card_of_mono1: -"A \ B \ |A| \o |B|" -using inj_on_id[of A] card_of_ordLeq[of A B] by fastforce - - -lemma card_of_mono2: -assumes "r \o r'" -shows "|Field r| \o |Field r'|" -proof- - obtain f where - 1: "well_order_on (Field r) r \ well_order_on (Field r) r \ embed r r' f" - using assms unfolding ordLeq_def - by (auto simp add: rel.well_order_on_Well_order) - hence "inj_on f (Field r) \ f ` (Field r) \ Field r'" - by (auto simp add: embed_inj_on embed_Field) - thus "|Field r| \o |Field r'|" using card_of_ordLeq by blast -qed - - -lemma card_of_cong: "r =o r' \ |Field r| =o |Field r'|" -by (simp add: ordIso_iff_ordLeq card_of_mono2) - - -lemma card_of_Field_ordLess: "Well_order r \ |Field r| \o r" -using card_of_least card_of_well_order_on rel.well_order_on_Well_order by blast - - -lemma card_of_Field_ordIso: -assumes "Card_order r" -shows "|Field r| =o r" -proof- - have "card_order_on (Field r) r" - using assms card_order_on_Card_order by blast - moreover have "card_order_on (Field r) |Field r|" - using card_of_card_order_on by blast - ultimately show ?thesis using card_order_on_ordIso by blast -qed - - -lemma Card_order_iff_ordIso_card_of: -"Card_order r = (r =o |Field r| )" -using ordIso_card_of_imp_Card_order card_of_Field_ordIso ordIso_symmetric by blast - - -lemma Card_order_iff_ordLeq_card_of: -"Card_order r = (r \o |Field r| )" -proof- - have "Card_order r = (r =o |Field r| )" - unfolding Card_order_iff_ordIso_card_of by simp - also have "... = (r \o |Field r| \ |Field r| \o r)" - unfolding ordIso_iff_ordLeq by simp - also have "... = (r \o |Field r| )" - using card_of_Field_ordLess - by (auto simp: card_of_Field_ordLess ordLeq_Well_order_simp) - finally show ?thesis . -qed - - -lemma Card_order_iff_Restr_underS: -assumes "Well_order r" -shows "Card_order r = (\a \ Field r. Restr r (rel.underS r a) o ?r'" - unfolding card_order_on_def by simp - moreover have "Field ?r' = ?A" - using 1 wo_rel.underS_ofilter Field_Restr_ofilter - unfolding wo_rel_def by fastforce - ultimately have "|?A| \o ?r'" by simp - also have "?r' o r" using card_of_least by blast - thus ?thesis using assms ordLeq_ordLess_trans by blast -qed - - -lemma internalize_card_of_ordLeq: -"( |A| \o r) = (\B \ Field r. |A| =o |B| \ |B| \o r)" -proof - assume "|A| \o r" - then obtain p where 1: "Field p \ Field r \ |A| =o p \ p \o r" - using internalize_ordLeq[of "|A|" r] by blast - hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast - hence "|Field p| =o p" using card_of_Field_ordIso by blast - hence "|A| =o |Field p| \ |Field p| \o r" - using 1 ordIso_equivalence ordIso_ordLeq_trans by blast - thus "\B \ Field r. |A| =o |B| \ |B| \o r" using 1 by blast -next - assume "\B \ Field r. |A| =o |B| \ |B| \o r" - thus "|A| \o r" using ordIso_ordLeq_trans by blast -qed - - -lemma internalize_card_of_ordLeq2: -"( |A| \o |C| ) = (\B \ C. |A| =o |B| \ |B| \o |C| )" -using internalize_card_of_ordLeq[of "A" "|C|"] Field_card_of[of C] by auto - - - -subsection {* Cardinals versus set operations on arbitrary sets *} - - -text{* Here we embark in a long journey of simple results showing -that the standard set-theoretic operations are well-behaved w.r.t. the notion of -cardinal -- essentially, this means that they preserve the ``cardinal identity" -@{text "=o"} and are monotonic w.r.t. @{text "\o"}. -*} - - -lemma card_of_empty: "|{}| \o |A|" -using card_of_ordLeq inj_on_id by blast - - -lemma card_of_empty1: -assumes "Well_order r \ Card_order r" -shows "|{}| \o r" -proof- - have "Well_order r" using assms unfolding card_order_on_def by auto - hence "|Field r| <=o r" - using assms card_of_Field_ordLess by blast - moreover have "|{}| \o |Field r|" by (simp add: card_of_empty) - ultimately show ?thesis using ordLeq_transitive by blast -qed - - -corollary Card_order_empty: -"Card_order r \ |{}| \o r" by (simp add: card_of_empty1) - - -lemma card_of_empty2: -assumes LEQ: "|A| =o |{}|" -shows "A = {}" -using assms card_of_ordIso[of A] bij_betw_empty2 by blast - - -lemma card_of_empty3: -assumes LEQ: "|A| \o |{}|" -shows "A = {}" -using assms -by (simp add: ordIso_iff_ordLeq card_of_empty1 card_of_empty2 - ordLeq_Well_order_simp) - - -lemma card_of_empty_ordIso: -"|{}::'a set| =o |{}::'b set|" -using card_of_ordIso unfolding bij_betw_def inj_on_def by blast - - -lemma card_of_image: -"|f ` A| <=o |A|" -proof(cases "A = {}", simp add: card_of_empty) - assume "A ~= {}" - hence "f ` A ~= {}" by auto - thus "|f ` A| \o |A|" - using card_of_ordLeq2[of "f ` A" A] by auto -qed - - -lemma surj_imp_ordLeq: -assumes "B <= f ` A" -shows "|B| <=o |A|" -proof- - have "|B| <=o |f ` A|" using assms card_of_mono1 by auto - thus ?thesis using card_of_image ordLeq_transitive by blast -qed - - -lemma card_of_ordLeqI2: -assumes "B \ f ` A" -shows "|B| \o |A|" -using assms by (metis surj_imp_ordLeq) - - -lemma card_of_singl_ordLeq: -assumes "A \ {}" -shows "|{b}| \o |A|" -proof- - obtain a where *: "a \ A" using assms by auto - let ?h = "\ b'::'b. if b' = b then a else undefined" - have "inj_on ?h {b} \ ?h ` {b} \ A" - using * unfolding inj_on_def by auto - thus ?thesis using card_of_ordLeq by blast -qed - - -corollary Card_order_singl_ordLeq: -"\Card_order r; Field r \ {}\ \ |{b}| \o r" -using card_of_singl_ordLeq[of "Field r" b] - card_of_Field_ordIso[of r] ordLeq_ordIso_trans by blast - - -lemma card_of_Pow: "|A| r o |A <+> B|" -proof- - have "Inl ` A \ A <+> B" by auto - thus ?thesis using inj_Inl[of A] card_of_ordLeq by blast -qed - - -corollary Card_order_Plus1: -"Card_order r \ r \o |(Field r) <+> B|" -using card_of_Plus1 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast - - -lemma card_of_Plus2: "|B| \o |A <+> B|" -proof- - have "Inr ` B \ A <+> B" by auto - thus ?thesis using inj_Inr[of B] card_of_ordLeq by blast -qed - - -corollary Card_order_Plus2: -"Card_order r \ r \o |A <+> (Field r)|" -using card_of_Plus2 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast - - -lemma card_of_Plus_empty1: "|A| =o |A <+> {}|" -proof- - have "bij_betw Inl A (A <+> {})" unfolding bij_betw_def inj_on_def by auto - thus ?thesis using card_of_ordIso by auto -qed - - -lemma card_of_Plus_empty2: "|A| =o |{} <+> A|" -proof- - have "bij_betw Inr A ({} <+> A)" unfolding bij_betw_def inj_on_def by auto - thus ?thesis using card_of_ordIso by auto -qed - - -lemma card_of_Plus_commute: "|A <+> B| =o |B <+> A|" -proof- - let ?f = "\(c::'a + 'b). case c of Inl a \ Inr a - | Inr b \ Inl b" - have "bij_betw ?f (A <+> B) (B <+> A)" - unfolding bij_betw_def inj_on_def by force - thus ?thesis using card_of_ordIso by blast -qed - - -lemma card_of_Plus_assoc: -fixes A :: "'a set" and B :: "'b set" and C :: "'c set" -shows "|(A <+> B) <+> C| =o |A <+> B <+> C|" -proof - - def f \ "\(k::('a + 'b) + 'c). - case k of Inl ab \ (case ab of Inl a \ Inl a - |Inr b \ Inr (Inl b)) - |Inr c \ Inr (Inr c)" - have "A <+> B <+> C \ f ` ((A <+> B) <+> C)" - proof - fix x assume x: "x \ A <+> B <+> C" - show "x \ f ` ((A <+> B) <+> C)" - proof(cases x) - case (Inl a) - hence "a \ A" "x = f (Inl (Inl a))" - using x unfolding f_def by auto - thus ?thesis by auto - next - case (Inr bc) note 1 = Inr show ?thesis - proof(cases bc) - case (Inl b) - hence "b \ B" "x = f (Inl (Inr b))" - using x 1 unfolding f_def by auto - thus ?thesis by auto - next - case (Inr c) - hence "c \ C" "x = f (Inr c)" - using x 1 unfolding f_def by auto - thus ?thesis by auto - qed - qed - qed - hence "bij_betw f ((A <+> B) <+> C) (A <+> B <+> C)" - unfolding bij_betw_def inj_on_def f_def by force - thus ?thesis using card_of_ordIso by blast -qed - - -lemma card_of_Plus_mono1: -assumes "|A| \o |B|" -shows "|A <+> C| \o |B <+> C|" -proof- - obtain f where 1: "inj_on f A \ f ` A \ B" - using assms card_of_ordLeq[of A] by fastforce - obtain g where g_def: - "g = (\d. case d of Inl a \ Inl(f a) | Inr (c::'c) \ Inr c)" by blast - have "inj_on g (A <+> C) \ g ` (A <+> C) \ (B <+> C)" - proof- - {fix d1 and d2 assume "d1 \ A <+> C \ d2 \ A <+> C" and - "g d1 = g d2" - hence "d1 = d2" using 1 unfolding inj_on_def - by(case_tac d1, case_tac d2, auto simp add: g_def) - } - moreover - {fix d assume "d \ A <+> C" - hence "g d \ B <+> C" using 1 - by(case_tac d, auto simp add: g_def) - } - ultimately show ?thesis unfolding inj_on_def by auto - qed - thus ?thesis using card_of_ordLeq by metis -qed - - -corollary ordLeq_Plus_mono1: -assumes "r \o r'" -shows "|(Field r) <+> C| \o |(Field r') <+> C|" -using assms card_of_mono2 card_of_Plus_mono1 by blast - - -lemma card_of_Plus_mono2: -assumes "|A| \o |B|" -shows "|C <+> A| \o |C <+> B|" -using assms card_of_Plus_mono1[of A B C] - card_of_Plus_commute[of C A] card_of_Plus_commute[of B C] - ordIso_ordLeq_trans[of "|C <+> A|"] ordLeq_ordIso_trans[of "|C <+> A|"] -by blast - - -corollary ordLeq_Plus_mono2: -assumes "r \o r'" -shows "|A <+> (Field r)| \o |A <+> (Field r')|" -using assms card_of_mono2 card_of_Plus_mono2 by blast - - -lemma card_of_Plus_mono: -assumes "|A| \o |B|" and "|C| \o |D|" -shows "|A <+> C| \o |B <+> D|" -using assms card_of_Plus_mono1[of A B C] card_of_Plus_mono2[of C D B] - ordLeq_transitive[of "|A <+> C|"] by blast - - -corollary ordLeq_Plus_mono: -assumes "r \o r'" and "p \o p'" -shows "|(Field r) <+> (Field p)| \o |(Field r') <+> (Field p')|" -using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Plus_mono by blast - - -lemma card_of_Plus_cong1: -assumes "|A| =o |B|" -shows "|A <+> C| =o |B <+> C|" -using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono1) - - -corollary ordIso_Plus_cong1: -assumes "r =o r'" -shows "|(Field r) <+> C| =o |(Field r') <+> C|" -using assms card_of_cong card_of_Plus_cong1 by blast - - -lemma card_of_Plus_cong2: -assumes "|A| =o |B|" -shows "|C <+> A| =o |C <+> B|" -using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono2) - - -corollary ordIso_Plus_cong2: -assumes "r =o r'" -shows "|A <+> (Field r)| =o |A <+> (Field r')|" -using assms card_of_cong card_of_Plus_cong2 by blast - - -lemma card_of_Plus_cong: -assumes "|A| =o |B|" and "|C| =o |D|" -shows "|A <+> C| =o |B <+> D|" -using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono) - - -corollary ordIso_Plus_cong: -assumes "r =o r'" and "p =o p'" -shows "|(Field r) <+> (Field p)| =o |(Field r') <+> (Field p')|" -using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Plus_cong by blast - - -lemma card_of_Un_Plus_ordLeq: -"|A \ B| \o |A <+> B|" -proof- - let ?f = "\ c. if c \ A then Inl c else Inr c" - have "inj_on ?f (A \ B) \ ?f ` (A \ B) \ A <+> B" - unfolding inj_on_def by auto - thus ?thesis using card_of_ordLeq by blast -qed - - -lemma card_of_Times1: -assumes "A \ {}" -shows "|B| \o |B \ A|" -proof(cases "B = {}", simp add: card_of_empty) - assume *: "B \ {}" - have "fst `(B \ A) = B" unfolding image_def using assms by auto - thus ?thesis using inj_on_iff_surj[of B "B \ A"] - card_of_ordLeq[of B "B \ A"] * by blast -qed - - -lemma card_of_Times_commute: "|A \ B| =o |B \ A|" -proof- - let ?f = "\(a::'a,b::'b). (b,a)" - have "bij_betw ?f (A \ B) (B \ A)" - unfolding bij_betw_def inj_on_def by auto - thus ?thesis using card_of_ordIso by blast -qed - - -lemma card_of_Times2: -assumes "A \ {}" shows "|B| \o |A \ B|" -using assms card_of_Times1[of A B] card_of_Times_commute[of B A] - ordLeq_ordIso_trans by blast - - -corollary Card_order_Times1: -"\Card_order r; B \ {}\ \ r \o |(Field r) \ B|" -using card_of_Times1[of B] card_of_Field_ordIso - ordIso_ordLeq_trans ordIso_symmetric by blast - - -corollary Card_order_Times2: -"\Card_order r; A \ {}\ \ r \o |A \ (Field r)|" -using card_of_Times2[of A] card_of_Field_ordIso - ordIso_ordLeq_trans ordIso_symmetric by blast - - -lemma card_of_Times3: "|A| \o |A \ A|" -using card_of_Times1[of A] -by(cases "A = {}", simp add: card_of_empty, blast) - - -lemma card_of_Plus_Times_bool: "|A <+> A| =o |A \ (UNIV::bool set)|" -proof- - let ?f = "\c::'a + 'a. case c of Inl a \ (a,True) - |Inr a \ (a,False)" - have "bij_betw ?f (A <+> A) (A \ (UNIV::bool set))" - proof- - {fix c1 and c2 assume "?f c1 = ?f c2" - hence "c1 = c2" - by(case_tac "c1", case_tac "c2", auto, case_tac "c2", auto) - } - moreover - {fix c assume "c \ A <+> A" - hence "?f c \ A \ (UNIV::bool set)" - by(case_tac c, auto) - } - moreover - {fix a bl assume *: "(a,bl) \ A \ (UNIV::bool set)" - have "(a,bl) \ ?f ` ( A <+> A)" - proof(cases bl) - assume bl hence "?f(Inl a) = (a,bl)" by auto - thus ?thesis using * by force - next - assume "\ bl" hence "?f(Inr a) = (a,bl)" by auto - thus ?thesis using * by force - qed - } - ultimately show ?thesis unfolding bij_betw_def inj_on_def by auto - qed - thus ?thesis using card_of_ordIso by blast -qed - - -lemma card_of_Times_mono1: -assumes "|A| \o |B|" -shows "|A \ C| \o |B \ C|" -proof- - obtain f where 1: "inj_on f A \ f ` A \ B" - using assms card_of_ordLeq[of A] by fastforce - obtain g where g_def: - "g = (\(a,c::'c). (f a,c))" by blast - have "inj_on g (A \ C) \ g ` (A \ C) \ (B \ C)" - using 1 unfolding inj_on_def using g_def by auto - thus ?thesis using card_of_ordLeq by metis -qed - - -corollary ordLeq_Times_mono1: -assumes "r \o r'" -shows "|(Field r) \ C| \o |(Field r') \ C|" -using assms card_of_mono2 card_of_Times_mono1 by blast - - -lemma card_of_Times_mono2: -assumes "|A| \o |B|" -shows "|C \ A| \o |C \ B|" -using assms card_of_Times_mono1[of A B C] - card_of_Times_commute[of C A] card_of_Times_commute[of B C] - ordIso_ordLeq_trans[of "|C \ A|"] ordLeq_ordIso_trans[of "|C \ A|"] -by blast - - -corollary ordLeq_Times_mono2: -assumes "r \o r'" -shows "|A \ (Field r)| \o |A \ (Field r')|" -using assms card_of_mono2 card_of_Times_mono2 by blast - - -lemma card_of_Sigma_mono1: -assumes "\i \ I. |A i| \o |B i|" -shows "|SIGMA i : I. A i| \o |SIGMA i : I. B i|" -proof- - have "\i. i \ I \ (\f. inj_on f (A i) \ f ` (A i) \ B i)" - using assms by (auto simp add: card_of_ordLeq) - with choice[of "\ i f. i \ I \ inj_on f (A i) \ f ` (A i) \ B i"] - obtain F where 1: "\i \ I. inj_on (F i) (A i) \ (F i) ` (A i) \ B i" by metis - obtain g where g_def: "g = (\(i,a::'b). (i,F i a))" by blast - have "inj_on g (Sigma I A) \ g ` (Sigma I A) \ (Sigma I B)" - using 1 unfolding inj_on_def using g_def by force - thus ?thesis using card_of_ordLeq by metis -qed - - -corollary card_of_Sigma_Times: -"\i \ I. |A i| \o |B| \ |SIGMA i : I. A i| \o |I \ B|" -using card_of_Sigma_mono1[of I A "\i. B"] . - - -lemma card_of_UNION_Sigma: -"|\i \ I. A i| \o |SIGMA i : I. A i|" -using Ex_inj_on_UNION_Sigma[of I A] card_of_ordLeq by metis - - -lemma card_of_bool: -assumes "a1 \ a2" -shows "|UNIV::bool set| =o |{a1,a2}|" -proof- - let ?f = "\ bl. case bl of True \ a1 | False \ a2" - have "bij_betw ?f UNIV {a1,a2}" - proof- - {fix bl1 and bl2 assume "?f bl1 = ?f bl2" - hence "bl1 = bl2" using assms by (case_tac bl1, case_tac bl2, auto) - } - moreover - {fix bl have "?f bl \ {a1,a2}" by (case_tac bl, auto) - } - moreover - {fix a assume *: "a \ {a1,a2}" - have "a \ ?f ` UNIV" - proof(cases "a = a1") - assume "a = a1" - hence "?f True = a" by auto thus ?thesis by blast - next - assume "a \ a1" hence "a = a2" using * by auto - hence "?f False = a" by auto thus ?thesis by blast - qed - } - ultimately show ?thesis unfolding bij_betw_def inj_on_def - by (metis image_subsetI order_eq_iff subsetI) - qed - thus ?thesis using card_of_ordIso by blast -qed - - -lemma card_of_Plus_Times_aux: -assumes A2: "a1 \ a2 \ {a1,a2} \ A" and - LEQ: "|A| \o |B|" -shows "|A <+> B| \o |A \ B|" -proof- - have 1: "|UNIV::bool set| \o |A|" - using A2 card_of_mono1[of "{a1,a2}"] card_of_bool[of a1 a2] - ordIso_ordLeq_trans[of "|UNIV::bool set|"] by metis - (* *) - have "|A <+> B| \o |B <+> B|" - using LEQ card_of_Plus_mono1 by blast - moreover have "|B <+> B| =o |B \ (UNIV::bool set)|" - using card_of_Plus_Times_bool by blast - moreover have "|B \ (UNIV::bool set)| \o |B \ A|" - using 1 by (simp add: card_of_Times_mono2) - moreover have " |B \ A| =o |A \ B|" - using card_of_Times_commute by blast - ultimately show "|A <+> B| \o |A \ B|" - using ordLeq_ordIso_trans[of "|A <+> B|" "|B <+> B|" "|B \ (UNIV::bool set)|"] - ordLeq_transitive[of "|A <+> B|" "|B \ (UNIV::bool set)|" "|B \ A|"] - ordLeq_ordIso_trans[of "|A <+> B|" "|B \ A|" "|A \ B|"] - by blast -qed - - -lemma card_of_Plus_Times: -assumes A2: "a1 \ a2 \ {a1,a2} \ A" and - B2: "b1 \ b2 \ {b1,b2} \ B" -shows "|A <+> B| \o |A \ B|" -proof- - {assume "|A| \o |B|" - hence ?thesis using assms by (auto simp add: card_of_Plus_Times_aux) - } - moreover - {assume "|B| \o |A|" - hence "|B <+> A| \o |B \ A|" - using assms by (auto simp add: card_of_Plus_Times_aux) - hence ?thesis - using card_of_Plus_commute card_of_Times_commute - ordIso_ordLeq_trans ordLeq_ordIso_trans by metis - } - ultimately show ?thesis - using card_of_Well_order[of A] card_of_Well_order[of B] - ordLeq_total[of "|A|"] by metis -qed - - -lemma card_of_ordLeq_finite: -assumes "|A| \o |B|" and "finite B" -shows "finite A" -using assms unfolding ordLeq_def -using embed_inj_on[of "|A|" "|B|"] embed_Field[of "|A|" "|B|"] - Field_card_of[of "A"] Field_card_of[of "B"] inj_on_finite[of _ "A" "B"] by fastforce - - -lemma card_of_ordLeq_infinite: -assumes "|A| \o |B|" and "infinite A" -shows "infinite B" -using assms card_of_ordLeq_finite by auto - - -lemma card_of_ordIso_finite: -assumes "|A| =o |B|" -shows "finite A = finite B" -using assms unfolding ordIso_def iso_def[abs_def] -by (auto simp: bij_betw_finite Field_card_of) - - -lemma card_of_ordIso_finite_Field: -assumes "Card_order r" and "r =o |A|" -shows "finite(Field r) = finite A" -using assms card_of_Field_ordIso card_of_ordIso_finite ordIso_equivalence by blast - - -subsection {* Cardinals versus set operations involving infinite sets *} - - -text{* Here we show that, for infinite sets, most set-theoretic constructions -do not increase the cardinality. The cornerstone for this is -theorem @{text "Card_order_Times_same_infinite"}, which states that self-product -does not increase cardinality -- the proof of this fact adapts a standard -set-theoretic argument, as presented, e.g., in the proof of theorem 1.5.11 -at page 47 in \cite{card-book}. Then everything else follows fairly easily. *} - - -lemma infinite_iff_card_of_nat: -"infinite A = ( |UNIV::nat set| \o |A| )" -by (auto simp add: infinite_iff_countable_subset card_of_ordLeq) - - -text{* The next two results correspond to the ZF fact that all infinite cardinals are -limit ordinals: *} - -lemma Card_order_infinite_not_under: -assumes CARD: "Card_order r" and INF: "infinite (Field r)" -shows "\ (\a. Field r = rel.under r a)" -proof(auto) - have 0: "Well_order r \ wo_rel r \ Refl r" - using CARD unfolding wo_rel_def card_order_on_def order_on_defs by auto - fix a assume *: "Field r = rel.under r a" - show False - proof(cases "a \ Field r") - assume Case1: "a \ Field r" - hence "rel.under r a = {}" unfolding Field_def rel.under_def by auto - thus False using INF * by auto - next - let ?r' = "Restr r (rel.underS r a)" - assume Case2: "a \ Field r" - hence 1: "rel.under r a = rel.underS r a \ {a} \ a \ rel.underS r a" - using 0 rel.Refl_under_underS rel.underS_notIn by fastforce - have 2: "wo_rel.ofilter r (rel.underS r a) \ rel.underS r a < Field r" - using 0 wo_rel.underS_ofilter * 1 Case2 by auto - hence "?r' Well_order ?r'" - using 2 0 Field_Restr_ofilter[of r] Well_order_Restr[of r] by blast - ultimately have "|rel.underS r a| f. bij_betw f (rel.under r a) (rel.underS r a)" - using infinite_imp_bij_betw[of "Field r" a] INF * 1 by auto - hence "|rel.under r a| =o |rel.underS r a|" using card_of_ordIso by blast - } - ultimately show False using not_ordLess_ordIso ordIso_symmetric by blast - qed -qed - - -lemma infinite_Card_order_limit: -assumes r: "Card_order r" and "infinite (Field r)" -and a: "a : Field r" -shows "EX b : Field r. a \ b \ (a,b) : r" -proof- - have "Field r \ rel.under r a" - using assms Card_order_infinite_not_under by blast - moreover have "rel.under r a \ Field r" - using rel.under_Field . - ultimately have "rel.under r a < Field r" by blast - then obtain b where 1: "b : Field r \ ~ (b,a) : r" - unfolding rel.under_def by blast - moreover have ba: "b \ a" - using 1 r unfolding card_order_on_def well_order_on_def - linear_order_on_def partial_order_on_def preorder_on_def refl_on_def by auto - ultimately have "(a,b) : r" - using a r unfolding card_order_on_def well_order_on_def linear_order_on_def - total_on_def by blast - thus ?thesis using 1 ba by auto -qed - - -theorem Card_order_Times_same_infinite: -assumes CO: "Card_order r" and INF: "infinite(Field r)" -shows "|Field r \ Field r| \o r" -proof- - obtain phi where phi_def: - "phi = (\r::'a rel. Card_order r \ infinite(Field r) \ - \ |Field r \ Field r| \o r )" by blast - have temp1: "\r. phi r \ Well_order r" - unfolding phi_def card_order_on_def by auto - have Ft: "\(\r. phi r)" - proof - assume "\r. phi r" - hence "{r. phi r} \ {} \ {r. phi r} \ {r. Well_order r}" - using temp1 by auto - then obtain r where 1: "phi r" and 2: "\r'. phi r' \ r \o r'" and - 3: "Card_order r \ Well_order r" - using exists_minim_Well_order[of "{r. phi r}"] temp1 phi_def by blast - let ?A = "Field r" let ?r' = "bsqr r" - have 4: "Well_order ?r' \ Field ?r' = ?A \ ?A \ |?A| =o r" - using 3 bsqr_Well_order Field_bsqr card_of_Field_ordIso by blast - have 5: "Card_order |?A \ ?A| \ Well_order |?A \ ?A|" - using card_of_Card_order card_of_Well_order by blast - (* *) - have "r ?A|" - using 1 3 5 ordLess_or_ordLeq unfolding phi_def by blast - moreover have "|?A \ ?A| \o ?r'" - using card_of_least[of "?A \ ?A"] 4 by auto - ultimately have "r bij_betw f ?A (?A \ ?A)" - unfolding ordLess_def embedS_def[abs_def] - by (auto simp add: Field_bsqr) - let ?B = "f ` ?A" - have "|?A| =o |?B|" - using 3 6 embed_inj_on inj_on_imp_bij_betw card_of_ordIso by blast - hence 8: "r =o |?B|" using 4 ordIso_transitive ordIso_symmetric by blast - (* *) - have "wo_rel.ofilter ?r' ?B" - using 6 embed_Field_ofilter 3 4 by blast - hence "wo_rel.ofilter ?r' ?B \ ?B \ ?A \ ?A \ ?B \ Field ?r'" - using 7 unfolding bij_betw_def using 6 3 embed_inj_on 4 by auto - hence temp2: "wo_rel.ofilter ?r' ?B \ ?B < ?A \ ?A" - using 4 wo_rel_def[of ?r'] wo_rel.ofilter_def[of ?r' ?B] by blast - have "\ (\a. Field r = rel.under r a)" - using 1 unfolding phi_def using Card_order_infinite_not_under[of r] by auto - then obtain A1 where temp3: "wo_rel.ofilter r A1 \ A1 < ?A" and 9: "?B \ A1 \ A1" - using temp2 3 bsqr_ofilter[of r ?B] by blast - hence "|?B| \o |A1 \ A1|" using card_of_mono1 by blast - hence 10: "r \o |A1 \ A1|" using 8 ordIso_ordLeq_trans by blast - let ?r1 = "Restr r A1" - have "?r1 o ?r1" using 3 Well_order_Restr card_of_least by blast - } - ultimately have 11: "|A1| Well_order |A1| \ Card_order |A1|" - using card_of_Card_order[of A1] card_of_Well_order[of A1] - by (simp add: Field_card_of) - moreover have "\ r \o | A1 |" - using temp4 11 3 using not_ordLeq_iff_ordLess by blast - ultimately have "infinite(Field |A1| ) \ Card_order |A1| \ \ r \o | A1 |" - by (simp add: card_of_card_order_on) - hence "|Field |A1| \ Field |A1| | \o |A1|" - using 2 unfolding phi_def by blast - hence "|A1 \ A1 | \o |A1|" using temp4 by auto - hence "r \o |A1|" using 10 ordLeq_transitive by blast - thus False using 11 not_ordLess_ordLeq by auto - qed - thus ?thesis using assms unfolding phi_def by blast -qed - - -corollary card_of_Times_same_infinite: -assumes "infinite A" -shows "|A \ A| =o |A|" -proof- - let ?r = "|A|" - have "Field ?r = A \ Card_order ?r" - using Field_card_of card_of_Card_order[of A] by fastforce - hence "|A \ A| \o |A|" - using Card_order_Times_same_infinite[of ?r] assms by auto - thus ?thesis using card_of_Times3 ordIso_iff_ordLeq by blast -qed - - -lemma card_of_Times_infinite: -assumes INF: "infinite A" and NE: "B \ {}" and LEQ: "|B| \o |A|" -shows "|A \ B| =o |A| \ |B \ A| =o |A|" -proof- - have "|A| \o |A \ B| \ |A| \o |B \ A|" - using assms by (simp add: card_of_Times1 card_of_Times2) - moreover - {have "|A \ B| \o |A \ A| \ |B \ A| \o |A \ A|" - using LEQ card_of_Times_mono1 card_of_Times_mono2 by blast - moreover have "|A \ A| =o |A|" using INF card_of_Times_same_infinite by blast - ultimately have "|A \ B| \o |A| \ |B \ A| \o |A|" - using ordLeq_ordIso_trans[of "|A \ B|"] ordLeq_ordIso_trans[of "|B \ A|"] by auto - } - ultimately show ?thesis by (simp add: ordIso_iff_ordLeq) -qed - - -corollary Card_order_Times_infinite: -assumes INF: "infinite(Field r)" and CARD: "Card_order r" and - NE: "Field p \ {}" and LEQ: "p \o r" -shows "| (Field r) \ (Field p) | =o r \ | (Field p) \ (Field r) | =o r" -proof- - have "|Field r \ Field p| =o |Field r| \ |Field p \ Field r| =o |Field r|" - using assms by (simp add: card_of_Times_infinite card_of_mono2) - thus ?thesis - using assms card_of_Field_ordIso[of r] - ordIso_transitive[of "|Field r \ Field p|"] - ordIso_transitive[of _ "|Field r|"] by blast -qed - - -lemma card_of_Sigma_ordLeq_infinite: -assumes INF: "infinite B" and - LEQ_I: "|I| \o |B|" and LEQ: "\i \ I. |A i| \o |B|" -shows "|SIGMA i : I. A i| \o |B|" -proof(cases "I = {}", simp add: card_of_empty) - assume *: "I \ {}" - have "|SIGMA i : I. A i| \o |I \ B|" - using LEQ card_of_Sigma_Times by blast - moreover have "|I \ B| =o |B|" - using INF * LEQ_I by (auto simp add: card_of_Times_infinite) - ultimately show ?thesis using ordLeq_ordIso_trans by blast -qed - - -lemma card_of_Sigma_ordLeq_infinite_Field: -assumes INF: "infinite (Field r)" and r: "Card_order r" and - LEQ_I: "|I| \o r" and LEQ: "\i \ I. |A i| \o r" -shows "|SIGMA i : I. A i| \o r" -proof- - let ?B = "Field r" - have 1: "r =o |?B| \ |?B| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast - hence "|I| \o |?B|" "\i \ I. |A i| \o |?B|" - using LEQ_I LEQ ordLeq_ordIso_trans by blast+ - hence "|SIGMA i : I. A i| \o |?B|" using INF LEQ - card_of_Sigma_ordLeq_infinite by blast - thus ?thesis using 1 ordLeq_ordIso_trans by blast -qed - - -lemma card_of_Times_ordLeq_infinite_Field: -"\infinite (Field r); |A| \o r; |B| \o r; Card_order r\ - \ |A <*> B| \o r" -by(simp add: card_of_Sigma_ordLeq_infinite_Field) - - -lemma card_of_Times_infinite_simps: -"\infinite A; B \ {}; |B| \o |A|\ \ |A \ B| =o |A|" -"\infinite A; B \ {}; |B| \o |A|\ \ |A| =o |A \ B|" -"\infinite A; B \ {}; |B| \o |A|\ \ |B \ A| =o |A|" -"\infinite A; B \ {}; |B| \o |A|\ \ |A| =o |B \ A|" -by (auto simp add: card_of_Times_infinite ordIso_symmetric) - - -lemma card_of_UNION_ordLeq_infinite: -assumes INF: "infinite B" and - LEQ_I: "|I| \o |B|" and LEQ: "\i \ I. |A i| \o |B|" -shows "|\ i \ I. A i| \o |B|" -proof(cases "I = {}", simp add: card_of_empty) - assume *: "I \ {}" - have "|\ i \ I. A i| \o |SIGMA i : I. A i|" - using card_of_UNION_Sigma by blast - moreover have "|SIGMA i : I. A i| \o |B|" - using assms card_of_Sigma_ordLeq_infinite by blast - ultimately show ?thesis using ordLeq_transitive by blast -qed - - -corollary card_of_UNION_ordLeq_infinite_Field: -assumes INF: "infinite (Field r)" and r: "Card_order r" and - LEQ_I: "|I| \o r" and LEQ: "\i \ I. |A i| \o r" -shows "|\ i \ I. A i| \o r" -proof- - let ?B = "Field r" - have 1: "r =o |?B| \ |?B| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast - hence "|I| \o |?B|" "\i \ I. |A i| \o |?B|" - using LEQ_I LEQ ordLeq_ordIso_trans by blast+ - hence "|\ i \ I. A i| \o |?B|" using INF LEQ - card_of_UNION_ordLeq_infinite by blast - thus ?thesis using 1 ordLeq_ordIso_trans by blast -qed - - -lemma card_of_Plus_infinite1: -assumes INF: "infinite A" and LEQ: "|B| \o |A|" -shows "|A <+> B| =o |A|" -proof(cases "B = {}", simp add: card_of_Plus_empty1 card_of_Plus_empty2 ordIso_symmetric) - let ?Inl = "Inl::'a \ 'a + 'b" let ?Inr = "Inr::'b \ 'a + 'b" - assume *: "B \ {}" - then obtain b1 where 1: "b1 \ B" by blast - show ?thesis - proof(cases "B = {b1}") - assume Case1: "B = {b1}" - have 2: "bij_betw ?Inl A ((?Inl ` A))" - unfolding bij_betw_def inj_on_def by auto - hence 3: "infinite (?Inl ` A)" - using INF bij_betw_finite[of ?Inl A] by blast - let ?A' = "?Inl ` A \ {?Inr b1}" - obtain g where "bij_betw g (?Inl ` A) ?A'" - using 3 infinite_imp_bij_betw2[of "?Inl ` A"] by auto - moreover have "?A' = A <+> B" using Case1 by blast - ultimately have "bij_betw g (?Inl ` A) (A <+> B)" by simp - hence "bij_betw (g o ?Inl) A (A <+> B)" - using 2 by (auto simp add: bij_betw_trans) - thus ?thesis using card_of_ordIso ordIso_symmetric by blast - next - assume Case2: "B \ {b1}" - with * 1 obtain b2 where 3: "b1 \ b2 \ {b1,b2} \ B" by fastforce - obtain f where "inj_on f B \ f ` B \ A" - using LEQ card_of_ordLeq[of B] by fastforce - with 3 have "f b1 \ f b2 \ {f b1, f b2} \ A" - unfolding inj_on_def by auto - with 3 have "|A <+> B| \o |A \ B|" - by (auto simp add: card_of_Plus_Times) - moreover have "|A \ B| =o |A|" - using assms * by (simp add: card_of_Times_infinite_simps) - ultimately have "|A <+> B| \o |A|" using ordLeq_ordIso_trans by metis - thus ?thesis using card_of_Plus1 ordIso_iff_ordLeq by blast - qed -qed - - -lemma card_of_Plus_infinite2: -assumes INF: "infinite A" and LEQ: "|B| \o |A|" -shows "|B <+> A| =o |A|" -using assms card_of_Plus_commute card_of_Plus_infinite1 -ordIso_equivalence by blast - - -lemma card_of_Plus_infinite: -assumes INF: "infinite A" and LEQ: "|B| \o |A|" -shows "|A <+> B| =o |A| \ |B <+> A| =o |A|" -using assms by (auto simp: card_of_Plus_infinite1 card_of_Plus_infinite2) - - -corollary Card_order_Plus_infinite: -assumes INF: "infinite(Field r)" and CARD: "Card_order r" and - LEQ: "p \o r" -shows "| (Field r) <+> (Field p) | =o r \ | (Field p) <+> (Field r) | =o r" -proof- - have "| Field r <+> Field p | =o | Field r | \ - | Field p <+> Field r | =o | Field r |" - using assms by (simp add: card_of_Plus_infinite card_of_mono2) - thus ?thesis - using assms card_of_Field_ordIso[of r] - ordIso_transitive[of "|Field r <+> Field p|"] - ordIso_transitive[of _ "|Field r|"] by blast -qed - - -subsection {* The cardinal $\omega$ and the finite cardinals *} - - -text{* The cardinal $\omega$, of natural numbers, shall be the standard non-strict -order relation on -@{text "nat"}, that we abbreviate by @{text "natLeq"}. The finite cardinals -shall be the restrictions of these relations to the numbers smaller than -fixed numbers @{text "n"}, that we abbreviate by @{text "natLeq_on n"}. *} - -abbreviation "(natLeq::(nat * nat) set) \ {(x,y). x \ y}" -abbreviation "(natLess::(nat * nat) set) \ {(x,y). x < y}" - -abbreviation natLeq_on :: "nat \ (nat * nat) set" -where "natLeq_on n \ {(x,y). x < n \ y < n \ x \ y}" - -lemma infinite_cartesian_product: -assumes "infinite A" "infinite B" -shows "infinite (A \ B)" -proof - assume "finite (A \ B)" - from assms(1) have "A \ {}" by auto - with `finite (A \ B)` have "finite B" using finite_cartesian_productD2 by auto - with assms(2) show False by simp -qed - - -subsubsection {* First as well-orders *} - - -lemma Field_natLeq: "Field natLeq = (UNIV::nat set)" -by(unfold Field_def, auto) - - -lemma natLeq_Refl: "Refl natLeq" -unfolding refl_on_def Field_def by auto - - -lemma natLeq_trans: "trans natLeq" -unfolding trans_def by auto - - -lemma natLeq_Preorder: "Preorder natLeq" -unfolding preorder_on_def -by (auto simp add: natLeq_Refl natLeq_trans) - - -lemma natLeq_antisym: "antisym natLeq" -unfolding antisym_def by auto - - -lemma natLeq_Partial_order: "Partial_order natLeq" -unfolding partial_order_on_def -by (auto simp add: natLeq_Preorder natLeq_antisym) - - -lemma natLeq_Total: "Total natLeq" -unfolding total_on_def by auto - - -lemma natLeq_Linear_order: "Linear_order natLeq" -unfolding linear_order_on_def -by (auto simp add: natLeq_Partial_order natLeq_Total) - - -lemma natLeq_natLess_Id: "natLess = natLeq - Id" -by auto - - -lemma natLeq_Well_order: "Well_order natLeq" -unfolding well_order_on_def -using natLeq_Linear_order wf_less natLeq_natLess_Id by auto - - -lemma closed_nat_set_iff: -assumes "\(m::nat) n. n \ A \ m \ n \ m \ A" -shows "A = UNIV \ (\n. A = {0 ..< n})" -proof- - {assume "A \ UNIV" hence "\n. n \ A" by blast - moreover obtain n where n_def: "n = (LEAST n. n \ A)" by blast - ultimately have 1: "n \ A \ (\m. m < n \ m \ A)" - using LeastI_ex[of "\ n. n \ A"] n_def Least_le[of "\ n. n \ A"] by fastforce - have "A = {0 ..< n}" - proof(auto simp add: 1) - fix m assume *: "m \ A" - {assume "n \ m" with assms * have "n \ A" by blast - hence False using 1 by auto - } - thus "m < n" by fastforce - qed - hence "\n. A = {0 ..< n}" by blast - } - thus ?thesis by blast -qed - - -lemma Field_natLeq_on: "Field (natLeq_on n) = {0 ..< n}" -unfolding Field_def by auto - - -lemma natLeq_underS_less: "rel.underS natLeq n = {0 ..< n}" -unfolding rel.underS_def by auto - - -lemma Restr_natLeq: "Restr natLeq {0 ..< n} = natLeq_on n" -by auto - - -lemma Restr_natLeq2: -"Restr natLeq (rel.underS natLeq n) = natLeq_on n" -by (auto simp add: Restr_natLeq natLeq_underS_less) - - -lemma natLeq_on_Well_order: "Well_order(natLeq_on n)" -using Restr_natLeq[of n] natLeq_Well_order - Well_order_Restr[of natLeq "{0.. m \ wo_rel.ofilter (natLeq_on m) {0 ..< n}" -by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def, - simp add: Field_natLeq_on, unfold rel.under_def, auto) - - -lemma natLeq_on_ofilter_iff: -"wo_rel.ofilter (natLeq_on m) A = (\n \ m. A = {0 ..< n})" -proof(rule iffI) - assume *: "wo_rel.ofilter (natLeq_on m) A" - hence 1: "A \ {0..n1 n2. n2 \ A \ n1 \ n2 \ n1 \ A" - using * by(fastforce simp add: natLeq_on_wo_rel wo_rel.ofilter_def rel.under_def) - hence "A = UNIV \ (\n. A = {0 ..< n})" using closed_nat_set_iff by blast - thus "\n \ m. A = {0 ..< n}" using 1 atLeastLessThan_less_eq by blast -next - assume "(\n\m. A = {0 ..< n})" - thus "wo_rel.ofilter (natLeq_on m) A" by (auto simp add: natLeq_on_ofilter_less_eq) -qed - - - -subsubsection {* Then as cardinals *} - - -lemma natLeq_Card_order: "Card_order natLeq" -proof(auto simp add: natLeq_Well_order - Card_order_iff_Restr_underS Restr_natLeq2, simp add: Field_natLeq) - fix n have "finite(Field (natLeq_on n))" - unfolding Field_natLeq_on by auto - moreover have "infinite(UNIV::nat set)" by auto - ultimately show "natLeq_on n o |A| )" -using infinite_iff_card_of_nat[of A] card_of_nat - ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast - - -corollary finite_iff_ordLess_natLeq: -"finite A = ( |A| finite A" -unfolding ordIso_def iso_def[abs_def] -by (auto simp: Field_natLeq_on bij_betw_finite Field_card_of) - - -lemma natLeq_on_Card_order: "Card_order (natLeq_on n)" -proof(unfold card_order_on_def, - auto simp add: natLeq_on_Well_order, simp add: Field_natLeq_on) - fix r assume "well_order_on {0..o r" - using finite_atLeastLessThan natLeq_on_well_order_on - finite_well_order_on_ordIso ordIso_iff_ordLeq by blast -qed - - -corollary card_of_Field_natLeq_on: -"|Field (natLeq_on n)| =o natLeq_on n" -using Field_natLeq_on natLeq_on_Card_order - Card_order_iff_ordIso_card_of[of "natLeq_on n"] - ordIso_symmetric[of "natLeq_on n"] by blast - - -corollary card_of_less: -"|{0 ..< n}| =o natLeq_on n" -using Field_natLeq_on card_of_Field_natLeq_on by auto - - -lemma natLeq_on_ordLeq_less_eq: -"((natLeq_on m) \o (natLeq_on n)) = (m \ n)" -proof - assume "natLeq_on m \o natLeq_on n" - then obtain f where "inj_on f {0.. f ` {0.. {0.. n" using atLeastLessThan_less_eq2 by blast -next - assume "m \ n" - hence "inj_on id {0.. id ` {0.. {0..o |{0..o natLeq_on n" - using card_of_less ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast -qed - - -lemma natLeq_on_ordLeq_less: -"((natLeq_on m) o |B| ) = (card A \ card B)" -using assms card_of_ordLeq[of A B] inj_on_iff_card_le[of A B] by blast - - -lemma finite_imp_card_of_natLeq_on: -assumes "finite A" -shows "|A| =o natLeq_on (card A)" -proof- - obtain h where "bij_betw h A {0 ..< card A}" - using assms ex_bij_betw_finite_nat by blast - thus ?thesis using card_of_ordIso card_of_less ordIso_equivalence by blast -qed - - -lemma finite_iff_card_of_natLeq_on: -"finite A = (\n. |A| =o natLeq_on n)" -using finite_imp_card_of_natLeq_on[of A] -by(auto simp add: ordIso_natLeq_on_imp_finite) - - - -subsection {* The successor of a cardinal *} - - -text{* First we define @{text "isCardSuc r r'"}, the notion of @{text "r'"} -being a successor cardinal of @{text "r"}. Although the definition does -not require @{text "r"} to be a cardinal, only this case will be meaningful. *} - - -definition isCardSuc :: "'a rel \ 'a set rel \ bool" -where -"isCardSuc r r' \ - Card_order r' \ r - (\(r''::'a set rel). Card_order r'' \ r r' \o r'')" - - -text{* Now we introduce the cardinal-successor operator @{text "cardSuc"}, -by picking {\em some} cardinal-order relation fulfilling @{text "isCardSuc"}. -Again, the picked item shall be proved unique up to order-isomorphism. *} - - -definition cardSuc :: "'a rel \ 'a set rel" -where -"cardSuc r \ SOME r'. isCardSuc r r'" - - -lemma exists_minim_Card_order: -"\R \ {}; \r \ R. Card_order r\ \ \r \ R. \r' \ R. r \o r'" -unfolding card_order_on_def using exists_minim_Well_order by blast - - -lemma exists_isCardSuc: -assumes "Card_order r" -shows "\r'. isCardSuc r r'" -proof- - let ?R = "{(r'::'a set rel). Card_order r' \ r ?R \ (\r \ ?R. Card_order r)" using assms - by (simp add: card_of_Card_order Card_order_Pow) - then obtain r where "r \ ?R \ (\r' \ ?R. r \o r')" - using exists_minim_Card_order[of ?R] by blast - thus ?thesis unfolding isCardSuc_def by auto -qed - - -lemma cardSuc_isCardSuc: -assumes "Card_order r" -shows "isCardSuc r (cardSuc r)" -unfolding cardSuc_def using assms -by (simp add: exists_isCardSuc someI_ex) - - -lemma cardSuc_Card_order: -"Card_order r \ Card_order(cardSuc r)" -using cardSuc_isCardSuc unfolding isCardSuc_def by blast - - -lemma cardSuc_greater: -"Card_order r \ r r \o cardSuc r" -using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast - - -text{* The minimality property of @{text "cardSuc"} originally present in its definition -is local to the type @{text "'a set rel"}, i.e., that of @{text "cardSuc r"}: *} - -lemma cardSuc_least_aux: -"\Card_order (r::'a rel); Card_order (r'::'a set rel); r \ cardSuc r \o r'" -using cardSuc_isCardSuc unfolding isCardSuc_def by blast - - -text{* But from this we can infer general minimality: *} - -lemma cardSuc_least: -assumes CARD: "Card_order r" and CARD': "Card_order r'" and LESS: "r o r'" -proof- - let ?p = "cardSuc r" - have 0: "Well_order ?p \ Well_order r'" - using assms cardSuc_Card_order unfolding card_order_on_def by blast - {assume "r' r'' o r''" using cardSuc_least_aux CARD by blast - hence False using 2 not_ordLess_ordLeq by blast - } - thus ?thesis using 0 ordLess_or_ordLeq by blast -qed - - -lemma cardSuc_ordLess_ordLeq: -assumes CARD: "Card_order r" and CARD': "Card_order r'" -shows "(r o r')" -proof(auto simp add: assms cardSuc_least) - assume "cardSuc r \o r'" - thus "r o r)" -proof- - have "Well_order r \ Well_order r'" - using assms unfolding card_order_on_def by auto - moreover have "Well_order(cardSuc r)" - using assms cardSuc_Card_order card_order_on_def by blast - ultimately show ?thesis - using assms cardSuc_ordLess_ordLeq[of r r'] - not_ordLeq_iff_ordLess[of r r'] not_ordLeq_iff_ordLess[of r' "cardSuc r"] by blast -qed - - -lemma cardSuc_mono_ordLeq: -assumes CARD: "Card_order r" and CARD': "Card_order r'" -shows "(cardSuc r \o cardSuc r') = (r \o r')" -using assms cardSuc_ordLeq_ordLess cardSuc_ordLess_ordLeq cardSuc_Card_order by blast - - -lemma cardSuc_invar_ordIso: -assumes CARD: "Card_order r" and CARD': "Card_order r'" -shows "(cardSuc r =o cardSuc r') = (r =o r')" -proof- - have 0: "Well_order r \ Well_order r' \ Well_order(cardSuc r) \ Well_order(cardSuc r')" - using assms by (simp add: card_order_on_well_order_on cardSuc_Card_order) - thus ?thesis - using ordIso_iff_ordLeq[of r r'] ordIso_iff_ordLeq - using cardSuc_mono_ordLeq[of r r'] cardSuc_mono_ordLeq[of r' r] assms by blast -qed - - -lemma cardSuc_natLeq_on_Suc: -"cardSuc(natLeq_on n) =o natLeq_on(Suc n)" -proof- - obtain r r' p where r_def: "r = natLeq_on n" and - r'_def: "r' = cardSuc(natLeq_on n)" and - p_def: "p = natLeq_on(Suc n)" by blast - (* Preliminary facts: *) - have CARD: "Card_order r \ Card_order r' \ Card_order p" unfolding r_def r'_def p_def - using cardSuc_ordLess_ordLeq natLeq_on_Card_order cardSuc_Card_order by blast - hence WELL: "Well_order r \ Well_order r' \ Well_order p" - unfolding card_order_on_def by force - have FIELD: "Field r = {0.. Field p = {0..<(Suc n)}" - unfolding r_def p_def Field_natLeq_on by simp - hence FIN: "finite (Field r)" by force - have "r o p" using CARD unfolding r_def r'_def p_def - using natLeq_on_ordLeq_less cardSuc_ordLess_ordLeq by blast - moreover have "p \o r'" - proof- - {assume "r' Suc n" and 3: "f ` (Field r') = {0.. n" using 2 by force - (* *) - have "bij_betw f (Field r') (f ` (Field r'))" - using 1 WELL embed_inj_on unfolding bij_betw_def by force - moreover have "finite(f ` (Field r'))" using 3 finite_atLeastLessThan[of 0 m] by force - ultimately have 5: "finite (Field r') \ card(Field r') = card (f ` (Field r'))" - using bij_betw_same_card bij_betw_finite by metis - hence "card(Field r') \ card(Field r)" using 3 4 FIELD by force - hence "|Field r'| \o |Field r|" using FIN 5 finite_card_of_iff_card2 by blast - hence False using LESS not_ordLess_ordLeq by auto - } - thus ?thesis using WELL CARD by (fastforce simp: not_ordLess_iff_ordLeq) - qed - ultimately show ?thesis using ordIso_iff_ordLeq unfolding r'_def p_def by blast -qed - - -lemma card_of_cardSuc_finite: -"finite(Field(cardSuc |A| )) = finite A" -proof - assume *: "finite (Field (cardSuc |A| ))" - have 0: "|Field(cardSuc |A| )| =o cardSuc |A|" - using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast - hence "|A| \o |Field(cardSuc |A| )|" - using card_of_Card_order[of A] cardSuc_ordLeq[of "|A|"] ordIso_symmetric - ordLeq_ordIso_trans by blast - thus "finite A" using * card_of_ordLeq_finite by blast -next - assume "finite A" - then obtain n where "|A| =o natLeq_on n" using finite_iff_card_of_natLeq_on by blast - hence "cardSuc |A| =o cardSuc(natLeq_on n)" - using card_of_Card_order cardSuc_invar_ordIso natLeq_on_Card_order by blast - hence "cardSuc |A| =o natLeq_on(Suc n)" - using cardSuc_natLeq_on_Suc ordIso_transitive by blast - hence "cardSuc |A| =o |{0..<(Suc n)}|" using card_of_less ordIso_equivalence by blast - moreover have "|Field (cardSuc |A| ) | =o cardSuc |A|" - using card_of_Field_ordIso cardSuc_Card_order card_of_Card_order by blast - ultimately have "|Field (cardSuc |A| ) | =o |{0..<(Suc n)}|" - using ordIso_equivalence by blast - thus "finite (Field (cardSuc |A| ))" - using card_of_ordIso_finite finite_atLeastLessThan by blast -qed - - -lemma cardSuc_finite: -assumes "Card_order r" -shows "finite (Field (cardSuc r)) = finite (Field r)" -proof- - let ?A = "Field r" - have "|?A| =o r" using assms by (simp add: card_of_Field_ordIso) - hence "cardSuc |?A| =o cardSuc r" using assms - by (simp add: card_of_Card_order cardSuc_invar_ordIso) - moreover have "|Field (cardSuc |?A| ) | =o cardSuc |?A|" - by (simp add: card_of_card_order_on Field_card_of card_of_Field_ordIso cardSuc_Card_order) - moreover - {have "|Field (cardSuc r) | =o cardSuc r" - using assms by (simp add: card_of_Field_ordIso cardSuc_Card_order) - hence "cardSuc r =o |Field (cardSuc r) |" - using ordIso_symmetric by blast - } - ultimately have "|Field (cardSuc |?A| ) | =o |Field (cardSuc r) |" - using ordIso_transitive by blast - hence "finite (Field (cardSuc |?A| )) = finite (Field (cardSuc r))" - using card_of_ordIso_finite by blast - thus ?thesis by (simp only: card_of_cardSuc_finite) -qed - - -lemma card_of_Plus_ordLess_infinite: -assumes INF: "infinite C" and - LESS1: "|A| B| B = {}") - assume Case1: "A = {} \ B = {}" - hence "|A| =o |A <+> B| \ |B| =o |A <+> B|" - using card_of_Plus_empty1 card_of_Plus_empty2 by blast - hence "|A <+> B| =o |A| \ |A <+> B| =o |B|" - using ordIso_symmetric[of "|A|"] ordIso_symmetric[of "|B|"] by blast - thus ?thesis using LESS1 LESS2 - ordIso_ordLess_trans[of "|A <+> B|" "|A|"] - ordIso_ordLess_trans[of "|A <+> B|" "|B|"] by blast -next - assume Case2: "\(A = {} \ B = {})" - {assume *: "|C| \o |A <+> B|" - hence "infinite (A <+> B)" using INF card_of_ordLeq_finite by blast - hence 1: "infinite A \ infinite B" using finite_Plus by blast - {assume Case21: "|A| \o |B|" - hence "infinite B" using 1 card_of_ordLeq_finite by blast - hence "|A <+> B| =o |B|" using Case2 Case21 - by (auto simp add: card_of_Plus_infinite) - hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast - } - moreover - {assume Case22: "|B| \o |A|" - hence "infinite A" using 1 card_of_ordLeq_finite by blast - hence "|A <+> B| =o |A|" using Case2 Case22 - by (auto simp add: card_of_Plus_infinite) - hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast - } - ultimately have False using ordLeq_total card_of_Well_order[of A] - card_of_Well_order[of B] by blast - } - thus ?thesis using ordLess_or_ordLeq[of "|A <+> B|" "|C|"] - card_of_Well_order[of "A <+> B"] card_of_Well_order[of "C"] by auto -qed - - -lemma card_of_Plus_ordLess_infinite_Field: -assumes INF: "infinite (Field r)" and r: "Card_order r" and - LESS1: "|A| B| |?C| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast - hence "|A| B| o r" and B: "|B| \o r" -and c: "Card_order r" -shows "|A <+> B| \o r" -proof- - let ?r' = "cardSuc r" - have "Card_order ?r' \ infinite (Field ?r')" using assms - by (simp add: cardSuc_Card_order cardSuc_finite) - moreover have "|A| B| o r" and B: "|B| \o r" -and "Card_order r" -shows "|A Un B| \o r" -using assms card_of_Plus_ordLeq_infinite_Field card_of_Un_Plus_ordLeq -ordLeq_transitive by blast - - - -subsection {* Regular cardinals *} - - -definition cofinal where -"cofinal A r \ - ALL a : Field r. EX b : A. a \ b \ (a,b) : r" - - -definition regular where -"regular r \ - ALL K. K \ Field r \ cofinal K r \ |K| =o r" - - -definition relChain where -"relChain r As \ - ALL i j. (i,j) \ r \ As i \ As j" - -lemma regular_UNION: -assumes r: "Card_order r" "regular r" -and As: "relChain r As" -and Bsub: "B \ (UN i : Field r. As i)" -and cardB: "|B| As i" -proof- - let ?phi = "%b j. j : Field r \ b : As j" - have "ALL b : B. EX j. ?phi b j" using Bsub by blast - then obtain f where f: "!! b. b : B \ ?phi b (f b)" - using bchoice[of B ?phi] by blast - let ?K = "f ` B" - {assume 1: "!! i. i : Field r \ ~ B \ As i" - have 2: "cofinal ?K r" - unfolding cofinal_def proof auto - fix i assume i: "i : Field r" - with 1 obtain b where b: "b : B \ b \ As i" by blast - hence "i \ f b \ ~ (f b,i) : r" - using As f unfolding relChain_def by auto - hence "i \ f b \ (i, f b) : r" using r - unfolding card_order_on_def well_order_on_def linear_order_on_def - total_on_def using i f b by auto - with b show "\b\B. i \ f b \ (i, f b) \ r" by blast - qed - moreover have "?K \ Field r" using f by blast - ultimately have "|?K| =o r" using 2 r unfolding regular_def by blast - moreover - { - have "|?K| <=o |B|" using card_of_image . - hence "|?K| (p \o r) = (p Field ?r'" and 2: "cofinal K ?r'" - hence "|K| \o |Field ?r'|" by (simp only: card_of_mono1) - also have 22: "|Field ?r'| =o ?r'" - using r' by (simp add: card_of_Field_ordIso[of ?r']) - finally have "|K| \o ?r'" . - moreover - {let ?L = "UN j : K. rel.underS ?r' j" - let ?J = "Field r" - have rJ: "r =o |?J|" - using r_card card_of_Field_ordIso ordIso_symmetric by blast - assume "|K| o |?J|" using rJ ordLeq_ordIso_trans by blast - moreover - {have "ALL j : K. |rel.underS ?r' j| o r" - using r' card_of_Card_order by blast - hence "ALL j : K. |rel.underS ?r' j| \o |?J|" - using rJ ordLeq_ordIso_trans by blast - } - ultimately have "|?L| \o |?J|" - using r_inf card_of_UNION_ordLeq_infinite by blast - hence "|?L| \o r" using rJ ordIso_symmetric ordLeq_ordIso_trans by blast - hence "|?L| ?L" - using 2 unfolding rel.underS_def cofinal_def by auto - hence "|Field ?r'| \o |?L|" by (simp add: card_of_mono1) - hence "?r' \o |?L|" - using 22 ordIso_ordLeq_trans ordIso_symmetric by blast - } - ultimately have "|?L| (UN i : Field (cardSuc r). As i)" -and cardB: "|B| <=o r" -shows "EX i : Field (cardSuc r). B \ As i" -proof- - let ?r' = "cardSuc r" - have "Card_order ?r' \ |B| a \ A. f a \ B) \ (\ a. a \ A \ f a = undefined)}" - -lemma Func_empty: -"Func {} B = {\x. undefined}" -unfolding Func_def by auto - -lemma Func_elim: -assumes "g \ Func A B" and "a \ A" -shows "\ b. b \ B \ g a = b" -using assms unfolding Func_def by (cases "g a = undefined") auto - -definition curr where -"curr A f \ \ a. if a \ A then \b. f (a,b) else undefined" - -lemma curr_in: -assumes f: "f \ Func (A <*> B) C" -shows "curr A f \ Func A (Func B C)" -using assms unfolding curr_def Func_def by auto - -lemma curr_inj: -assumes "f1 \ Func (A <*> B) C" and "f2 \ Func (A <*> B) C" -shows "curr A f1 = curr A f2 \ f1 = f2" -proof safe - assume c: "curr A f1 = curr A f2" - show "f1 = f2" - proof (rule ext, clarify) - fix a b show "f1 (a, b) = f2 (a, b)" - proof (cases "(a,b) \ A <*> B") - case False - thus ?thesis using assms unfolding Func_def by auto - next - case True hence a: "a \ A" and b: "b \ B" by auto - thus ?thesis - using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp - qed - qed -qed - -lemma curr_surj: -assumes "g \ Func A (Func B C)" -shows "\ f \ Func (A <*> B) C. curr A f = g" -proof - let ?f = "\ ab. if fst ab \ A \ snd ab \ B then g (fst ab) (snd ab) else undefined" - show "curr A ?f = g" - proof (rule ext) - fix a show "curr A ?f a = g a" - proof (cases "a \ A") - case False - hence "g a = undefined" using assms unfolding Func_def by auto - thus ?thesis unfolding curr_def using False by simp - next - case True - obtain g1 where "g1 \ Func B C" and "g a = g1" - using assms using Func_elim[OF assms True] by blast - thus ?thesis using True unfolding Func_def curr_def by auto - qed - qed - show "?f \ Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto -qed - -lemma bij_betw_curr: -"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))" -unfolding bij_betw_def inj_on_def image_def -using curr_in curr_inj curr_surj by blast - -lemma card_of_Func_Times: -"|Func (A <*> B) C| =o |Func A (Func B C)|" -unfolding card_of_ordIso[symmetric] -using bij_betw_curr by blast - -definition Func_map where -"Func_map B2 f1 f2 g b2 \ if b2 \ B2 then f1 (g (f2 b2)) else undefined" - -lemma Func_map: -assumes g: "g \ Func A2 A1" and f1: "f1 ` A1 \ B1" and f2: "f2 ` B2 \ A2" -shows "Func_map B2 f1 f2 g \ Func B2 B1" -using assms unfolding Func_def Func_map_def mem_Collect_eq by auto - -lemma Func_non_emp: -assumes "B \ {}" -shows "Func A B \ {}" -proof- - obtain b where b: "b \ B" using assms by auto - hence "(\ a. if a \ A then b else undefined) \ Func A B" unfolding Func_def by auto - thus ?thesis by blast -qed - -lemma Func_is_emp: -"Func A B = {} \ A \ {} \ B = {}" (is "?L \ ?R") -proof - assume L: ?L - moreover {assume "A = {}" hence False using L Func_empty by auto} - moreover {assume "B \ {}" hence False using L Func_non_emp by metis} - ultimately show ?R by blast -next - assume R: ?R - moreover - {fix f assume "f \ Func A B" - moreover obtain a where "a \ A" using R by blast - ultimately obtain b where "b \ B" unfolding Func_def by(cases "f a = undefined", force+) - with R have False by auto - } - thus ?L by blast -qed - -lemma Func_map_surj: -assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \ A2" -and B2A2: "B2 = {} \ A2 = {}" -shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1" -proof(cases "B2 = {}") - case True - thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def) -next - case False note B2 = False - show ?thesis - proof safe - fix h assume h: "h \ Func B2 B1" - def j1 \ "inv_into A1 f1" - have "\ a2 \ f2 ` B2. \ b2. b2 \ B2 \ f2 b2 = a2" by blast - then obtain k where k: "\ a2 \ f2 ` B2. k a2 \ B2 \ f2 (k a2) = a2" by metis - {fix b2 assume b2: "b2 \ B2" - hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto - moreover have "k (f2 b2) \ B2" using b2 A2(2) k by auto - ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast - } note kk = this - obtain b22 where b22: "b22 \ B2" using B2 by auto - def j2 \ "\ a2. if a2 \ f2 ` B2 then k a2 else b22" - have j2A2: "j2 ` A2 \ B2" unfolding j2_def using k b22 by auto - have j2: "\ b2. b2 \ B2 \ j2 (f2 b2) = b2" - using kk unfolding j2_def by auto - def g \ "Func_map A2 j1 j2 h" - have "Func_map B2 f1 f2 g = h" - proof (rule ext) - fix b2 show "Func_map B2 f1 f2 g b2 = h b2" - proof(cases "b2 \ B2") - case True - show ?thesis - proof (cases "h b2 = undefined") - case True - hence b1: "h b2 \ f1 ` A1" using h `b2 \ B2` unfolding B1 Func_def by auto - show ?thesis using A2 f_inv_into_f[OF b1] - unfolding True g_def Func_map_def j1_def j2[OF `b2 \ B2`] by auto - qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def, - auto intro: f_inv_into_f) - qed(insert h, unfold Func_def Func_map_def, auto) - qed - moreover have "g \ Func A2 A1" unfolding g_def apply(rule Func_map[OF h]) - using inv_into_into j2A2 B1 A2 inv_into_into - unfolding j1_def image_def by fast+ - ultimately show "h \ Func_map B2 f1 f2 ` Func A2 A1" - unfolding Func_map_def[abs_def] unfolding image_def by auto - qed(insert B1 Func_map[OF _ _ A2(2)], auto) -qed - -lemma card_of_Pow_Func: -"|Pow A| =o |Func A (UNIV::bool set)|" -proof- - def F \ "\ A' a. if a \ A then (if a \ A' then True else False) - else undefined" - have "bij_betw F (Pow A) (Func A (UNIV::bool set))" - unfolding bij_betw_def inj_on_def proof (intro ballI impI conjI) - fix A1 A2 assume "A1 \ Pow A" "A2 \ Pow A" "F A1 = F A2" - thus "A1 = A2" unfolding F_def Pow_def fun_eq_iff by (auto split: split_if_asm) - next - show "F ` Pow A = Func A UNIV" - proof safe - fix f assume f: "f \ Func A (UNIV::bool set)" - show "f \ F ` Pow A" unfolding image_def mem_Collect_eq proof(intro bexI) - let ?A1 = "{a \ A. f a = True}" - show "f = F ?A1" unfolding F_def apply(rule ext) - using f unfolding Func_def mem_Collect_eq by auto - qed auto - qed(unfold Func_def mem_Collect_eq F_def, auto) - qed - thus ?thesis unfolding card_of_ordIso[symmetric] by blast -qed - -lemma card_of_Func_UNIV: -"|Func (UNIV::'a set) (B::'b set)| =o |{f::'a \ 'b. range f \ B}|" -apply(rule ordIso_symmetric) proof(intro card_of_ordIsoI) - let ?F = "\ f (a::'a). ((f a)::'b)" - show "bij_betw ?F {f. range f \ B} (Func UNIV B)" - unfolding bij_betw_def inj_on_def proof safe - fix h :: "'a \ 'b" assume h: "h \ Func UNIV B" - hence "\ a. \ b. h a = b" unfolding Func_def by auto - then obtain f where f: "\ a. h a = f a" by metis - hence "range f \ B" using h unfolding Func_def by auto - thus "h \ (\f a. f a) ` {f. range f \ B}" using f unfolding image_def by auto - qed(unfold Func_def fun_eq_iff, auto) -qed - -end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Constructions_on_Wellorders.thy --- a/src/HOL/Cardinals/Constructions_on_Wellorders.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/Constructions_on_Wellorders.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* Constructions on Wellorders *} theory Constructions_on_Wellorders -imports Constructions_on_Wellorders_LFP Wellorder_Embedding +imports Constructions_on_Wellorders_FP Wellorder_Embedding begin declare diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Constructions_on_Wellorders_FP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Cardinals/Constructions_on_Wellorders_FP.thy Mon Nov 18 18:04:45 2013 +0100 @@ -0,0 +1,1621 @@ +(* Title: HOL/Cardinals/Constructions_on_Wellorders_FP.thy + Author: Andrei Popescu, TU Muenchen + Copyright 2012 + +Constructions on wellorders (FP). +*) + +header {* Constructions on Wellorders (FP) *} + +theory Constructions_on_Wellorders_FP +imports Wellorder_Embedding_FP +begin + + +text {* In this section, we study basic constructions on well-orders, such as restriction to +a set/order filter, copy via direct images, ordinal-like sum of disjoint well-orders, +and bounded square. We also define between well-orders +the relations @{text "ordLeq"}, of being embedded (abbreviated @{text "\o"}), +@{text "ordLess"}, of being strictly embedded (abbreviated @{text " 'a set \ 'a rel" +where "Restr r A \ r Int (A \ A)" + + +lemma Restr_subset: +"A \ B \ Restr (Restr r B) A = Restr r A" +by blast + + +lemma Restr_Field: "Restr r (Field r) = r" +unfolding Field_def by auto + + +lemma Refl_Restr: "Refl r \ Refl(Restr r A)" +unfolding refl_on_def Field_def by auto + + +lemma antisym_Restr: +"antisym r \ antisym(Restr r A)" +unfolding antisym_def Field_def by auto + + +lemma Total_Restr: +"Total r \ Total(Restr r A)" +unfolding total_on_def Field_def by auto + + +lemma trans_Restr: +"trans r \ trans(Restr r A)" +unfolding trans_def Field_def by blast + + +lemma Preorder_Restr: +"Preorder r \ Preorder(Restr r A)" +unfolding preorder_on_def by (simp add: Refl_Restr trans_Restr) + + +lemma Partial_order_Restr: +"Partial_order r \ Partial_order(Restr r A)" +unfolding partial_order_on_def by (simp add: Preorder_Restr antisym_Restr) + + +lemma Linear_order_Restr: +"Linear_order r \ Linear_order(Restr r A)" +unfolding linear_order_on_def by (simp add: Partial_order_Restr Total_Restr) + + +lemma Well_order_Restr: +assumes "Well_order r" +shows "Well_order(Restr r A)" +proof- + have "Restr r A - Id \ r - Id" using Restr_subset by blast + hence "wf(Restr r A - Id)" using assms + using well_order_on_def wf_subset by blast + thus ?thesis using assms unfolding well_order_on_def + by (simp add: Linear_order_Restr) +qed + + +lemma Field_Restr_subset: "Field(Restr r A) \ A" +by (auto simp add: Field_def) + + +lemma Refl_Field_Restr: +"Refl r \ Field(Restr r A) = (Field r) Int A" +by (auto simp add: refl_on_def Field_def) + + +lemma Refl_Field_Restr2: +"\Refl r; A \ Field r\ \ Field(Restr r A) = A" +by (auto simp add: Refl_Field_Restr) + + +lemma well_order_on_Restr: +assumes WELL: "Well_order r" and SUB: "A \ Field r" +shows "well_order_on A (Restr r A)" +using assms +using Well_order_Restr[of r A] Refl_Field_Restr2[of r A] + order_on_defs[of "Field r" r] by auto + + +subsection {* Order filters versus restrictions and embeddings *} + + +lemma Field_Restr_ofilter: +"\Well_order r; wo_rel.ofilter r A\ \ Field(Restr r A) = A" +by (auto simp add: wo_rel_def wo_rel.ofilter_def wo_rel.REFL Refl_Field_Restr2) + + +lemma ofilter_Restr_under: +assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and IN: "a \ A" +shows "rel.under (Restr r A) a = rel.under r a" +using assms wo_rel_def +proof(auto simp add: wo_rel.ofilter_def rel.under_def) + fix b assume *: "a \ A" and "(b,a) \ r" + hence "b \ rel.under r a \ a \ Field r" + unfolding rel.under_def using Field_def by fastforce + thus "b \ A" using * assms by (auto simp add: wo_rel_def wo_rel.ofilter_def) +qed + + +lemma ofilter_embed: +assumes "Well_order r" +shows "wo_rel.ofilter r A = (A \ Field r \ embed (Restr r A) r id)" +proof + assume *: "wo_rel.ofilter r A" + show "A \ Field r \ embed (Restr r A) r id" + proof(unfold embed_def, auto) + fix a assume "a \ A" thus "a \ Field r" using assms * + by (auto simp add: wo_rel_def wo_rel.ofilter_def) + next + fix a assume "a \ Field (Restr r A)" + thus "bij_betw id (rel.under (Restr r A) a) (rel.under r a)" using assms * + by (simp add: ofilter_Restr_under Field_Restr_ofilter) + qed +next + assume *: "A \ Field r \ embed (Restr r A) r id" + hence "Field(Restr r A) \ Field r" + using assms embed_Field[of "Restr r A" r id] id_def + Well_order_Restr[of r] by auto + {fix a assume "a \ A" + hence "a \ Field(Restr r A)" using * assms + by (simp add: order_on_defs Refl_Field_Restr2) + hence "bij_betw id (rel.under (Restr r A) a) (rel.under r a)" + using * unfolding embed_def by auto + hence "rel.under r a \ rel.under (Restr r A) a" + unfolding bij_betw_def by auto + also have "\ \ Field(Restr r A)" by (simp add: rel.under_Field) + also have "\ \ A" by (simp add: Field_Restr_subset) + finally have "rel.under r a \ A" . + } + thus "wo_rel.ofilter r A" using assms * by (simp add: wo_rel_def wo_rel.ofilter_def) +qed + + +lemma ofilter_Restr_Int: +assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" +shows "wo_rel.ofilter (Restr r B) (A Int B)" +proof- + let ?rB = "Restr r B" + have Well: "wo_rel r" unfolding wo_rel_def using WELL . + hence Refl: "Refl r" by (simp add: wo_rel.REFL) + hence Field: "Field ?rB = Field r Int B" + using Refl_Field_Restr by blast + have WellB: "wo_rel ?rB \ Well_order ?rB" using WELL + by (simp add: Well_order_Restr wo_rel_def) + (* Main proof *) + show ?thesis using WellB assms + proof(auto simp add: wo_rel.ofilter_def rel.under_def) + fix a assume "a \ A" and *: "a \ B" + hence "a \ Field r" using OFA Well by (auto simp add: wo_rel.ofilter_def) + with * show "a \ Field ?rB" using Field by auto + next + fix a b assume "a \ A" and "(b,a) \ r" + thus "b \ A" using Well OFA by (auto simp add: wo_rel.ofilter_def rel.under_def) + qed +qed + + +lemma ofilter_Restr_subset: +assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and SUB: "A \ B" +shows "wo_rel.ofilter (Restr r B) A" +proof- + have "A Int B = A" using SUB by blast + thus ?thesis using assms ofilter_Restr_Int[of r A B] by auto +qed + + +lemma ofilter_subset_embed: +assumes WELL: "Well_order r" and + OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" +shows "(A \ B) = (embed (Restr r A) (Restr r B) id)" +proof- + let ?rA = "Restr r A" let ?rB = "Restr r B" + have Well: "wo_rel r" unfolding wo_rel_def using WELL . + hence Refl: "Refl r" by (simp add: wo_rel.REFL) + hence FieldA: "Field ?rA = Field r Int A" + using Refl_Field_Restr by blast + have FieldB: "Field ?rB = Field r Int B" + using Refl Refl_Field_Restr by blast + have WellA: "wo_rel ?rA \ Well_order ?rA" using WELL + by (simp add: Well_order_Restr wo_rel_def) + have WellB: "wo_rel ?rB \ Well_order ?rB" using WELL + by (simp add: Well_order_Restr wo_rel_def) + (* Main proof *) + show ?thesis + proof + assume *: "A \ B" + hence "wo_rel.ofilter (Restr r B) A" using assms + by (simp add: ofilter_Restr_subset) + hence "embed (Restr ?rB A) (Restr r B) id" + using WellB ofilter_embed[of "?rB" A] by auto + thus "embed (Restr r A) (Restr r B) id" + using * by (simp add: Restr_subset) + next + assume *: "embed (Restr r A) (Restr r B) id" + {fix a assume **: "a \ A" + hence "a \ Field r" using Well OFA by (auto simp add: wo_rel.ofilter_def) + with ** FieldA have "a \ Field ?rA" by auto + hence "a \ Field ?rB" using * WellA embed_Field[of ?rA ?rB id] by auto + hence "a \ B" using FieldB by auto + } + thus "A \ B" by blast + qed +qed + + +lemma ofilter_subset_embedS_iso: +assumes WELL: "Well_order r" and + OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" +shows "((A < B) = (embedS (Restr r A) (Restr r B) id)) \ + ((A = B) = (iso (Restr r A) (Restr r B) id))" +proof- + let ?rA = "Restr r A" let ?rB = "Restr r B" + have Well: "wo_rel r" unfolding wo_rel_def using WELL . + hence Refl: "Refl r" by (simp add: wo_rel.REFL) + hence "Field ?rA = Field r Int A" + using Refl_Field_Restr by blast + hence FieldA: "Field ?rA = A" using OFA Well + by (auto simp add: wo_rel.ofilter_def) + have "Field ?rB = Field r Int B" + using Refl Refl_Field_Restr by blast + hence FieldB: "Field ?rB = B" using OFB Well + by (auto simp add: wo_rel.ofilter_def) + (* Main proof *) + show ?thesis unfolding embedS_def iso_def + using assms ofilter_subset_embed[of r A B] + FieldA FieldB bij_betw_id_iff[of A B] by auto +qed + + +lemma ofilter_subset_embedS: +assumes WELL: "Well_order r" and + OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" +shows "(A < B) = embedS (Restr r A) (Restr r B) id" +using assms by (simp add: ofilter_subset_embedS_iso) + + +lemma embed_implies_iso_Restr: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r' r f" +shows "iso r' (Restr r (f ` (Field r'))) f" +proof- + let ?A' = "Field r'" + let ?r'' = "Restr r (f ` ?A')" + have 0: "Well_order ?r''" using WELL Well_order_Restr by blast + have 1: "wo_rel.ofilter r (f ` ?A')" using assms embed_Field_ofilter by blast + hence "Field ?r'' = f ` (Field r')" using WELL Field_Restr_ofilter by blast + hence "bij_betw f ?A' (Field ?r'')" + using EMB embed_inj_on WELL' unfolding bij_betw_def by blast + moreover + {have "\a b. (a,b) \ r' \ a \ Field r' \ b \ Field r'" + unfolding Field_def by auto + hence "compat r' ?r'' f" + using assms embed_iff_compat_inj_on_ofilter + unfolding compat_def by blast + } + ultimately show ?thesis using WELL' 0 iso_iff3 by blast +qed + + +subsection {* The strict inclusion on proper ofilters is well-founded *} + + +definition ofilterIncl :: "'a rel \ 'a set rel" +where +"ofilterIncl r \ {(A,B). wo_rel.ofilter r A \ A \ Field r \ + wo_rel.ofilter r B \ B \ Field r \ A < B}" + + +lemma wf_ofilterIncl: +assumes WELL: "Well_order r" +shows "wf(ofilterIncl r)" +proof- + have Well: "wo_rel r" using WELL by (simp add: wo_rel_def) + hence Lo: "Linear_order r" by (simp add: wo_rel.LIN) + let ?h = "(\ A. wo_rel.suc r A)" + let ?rS = "r - Id" + have "wf ?rS" using WELL by (simp add: order_on_defs) + moreover + have "compat (ofilterIncl r) ?rS ?h" + proof(unfold compat_def ofilterIncl_def, + intro allI impI, simp, elim conjE) + fix A B + assume *: "wo_rel.ofilter r A" "A \ Field r" and + **: "wo_rel.ofilter r B" "B \ Field r" and ***: "A < B" + then obtain a and b where 0: "a \ Field r \ b \ Field r" and + 1: "A = rel.underS r a \ B = rel.underS r b" + using Well by (auto simp add: wo_rel.ofilter_underS_Field) + hence "a \ b" using *** by auto + moreover + have "(a,b) \ r" using 0 1 Lo *** + by (auto simp add: rel.underS_incl_iff) + moreover + have "a = wo_rel.suc r A \ b = wo_rel.suc r B" + using Well 0 1 by (simp add: wo_rel.suc_underS) + ultimately + show "(wo_rel.suc r A, wo_rel.suc r B) \ r \ wo_rel.suc r A \ wo_rel.suc r B" + by simp + qed + ultimately show "wf (ofilterIncl r)" by (simp add: compat_wf) +qed + + + +subsection {* Ordering the well-orders by existence of embeddings *} + + +text {* We define three relations between well-orders: +\begin{itemize} +\item @{text "ordLeq"}, of being embedded (abbreviated @{text "\o"}); +\item @{text "ordLess"}, of being strictly embedded (abbreviated @{text ""}, @{text "<"}, @{text "="} associated to a total order on a set. +*} + + +definition ordLeq :: "('a rel * 'a' rel) set" +where +"ordLeq = {(r,r'). Well_order r \ Well_order r' \ (\f. embed r r' f)}" + + +abbreviation ordLeq2 :: "'a rel \ 'a' rel \ bool" (infix "<=o" 50) +where "r <=o r' \ (r,r') \ ordLeq" + + +abbreviation ordLeq3 :: "'a rel \ 'a' rel \ bool" (infix "\o" 50) +where "r \o r' \ r <=o r'" + + +definition ordLess :: "('a rel * 'a' rel) set" +where +"ordLess = {(r,r'). Well_order r \ Well_order r' \ (\f. embedS r r' f)}" + +abbreviation ordLess2 :: "'a rel \ 'a' rel \ bool" (infix " (r,r') \ ordLess" + + +definition ordIso :: "('a rel * 'a' rel) set" +where +"ordIso = {(r,r'). Well_order r \ Well_order r' \ (\f. iso r r' f)}" + +abbreviation ordIso2 :: "'a rel \ 'a' rel \ bool" (infix "=o" 50) +where "r =o r' \ (r,r') \ ordIso" + + +lemmas ordRels_def = ordLeq_def ordLess_def ordIso_def + +lemma ordLeq_Well_order_simp: +assumes "r \o r'" +shows "Well_order r \ Well_order r'" +using assms unfolding ordLeq_def by simp + + +text{* Notice that the relations @{text "\o"}, @{text " r \o r" +unfolding ordLeq_def using id_embed[of r] by blast + + +lemma ordLeq_transitive[trans]: +assumes *: "r \o r'" and **: "r' \o r''" +shows "r \o r''" +proof- + obtain f and f' + where 1: "Well_order r \ Well_order r' \ Well_order r''" and + "embed r r' f" and "embed r' r'' f'" + using * ** unfolding ordLeq_def by blast + hence "embed r r'' (f' o f)" + using comp_embed[of r r' f r'' f'] by auto + thus "r \o r''" unfolding ordLeq_def using 1 by auto +qed + + +lemma ordLeq_total: +"\Well_order r; Well_order r'\ \ r \o r' \ r' \o r" +unfolding ordLeq_def using wellorders_totally_ordered by blast + + +lemma ordIso_reflexive: +"Well_order r \ r =o r" +unfolding ordIso_def using id_iso[of r] by blast + + +lemma ordIso_transitive[trans]: +assumes *: "r =o r'" and **: "r' =o r''" +shows "r =o r''" +proof- + obtain f and f' + where 1: "Well_order r \ Well_order r' \ Well_order r''" and + "iso r r' f" and 3: "iso r' r'' f'" + using * ** unfolding ordIso_def by auto + hence "iso r r'' (f' o f)" + using comp_iso[of r r' f r'' f'] by auto + thus "r =o r''" unfolding ordIso_def using 1 by auto +qed + + +lemma ordIso_symmetric: +assumes *: "r =o r'" +shows "r' =o r" +proof- + obtain f where 1: "Well_order r \ Well_order r'" and + 2: "embed r r' f \ bij_betw f (Field r) (Field r')" + using * by (auto simp add: ordIso_def iso_def) + let ?f' = "inv_into (Field r) f" + have "embed r' r ?f' \ bij_betw ?f' (Field r') (Field r)" + using 1 2 by (simp add: bij_betw_inv_into inv_into_Field_embed_bij_betw) + thus "r' =o r" unfolding ordIso_def using 1 by (auto simp add: iso_def) +qed + + +lemma ordLeq_ordLess_trans[trans]: +assumes "r \o r'" and " r' Well_order r''" + using assms unfolding ordLeq_def ordLess_def by auto + thus ?thesis using assms unfolding ordLeq_def ordLess_def + using embed_comp_embedS by blast +qed + + +lemma ordLess_ordLeq_trans[trans]: +assumes "r o r''" +shows "r Well_order r''" + using assms unfolding ordLeq_def ordLess_def by auto + thus ?thesis using assms unfolding ordLeq_def ordLess_def + using embedS_comp_embed by blast +qed + + +lemma ordLeq_ordIso_trans[trans]: +assumes "r \o r'" and " r' =o r''" +shows "r \o r''" +proof- + have "Well_order r \ Well_order r''" + using assms unfolding ordLeq_def ordIso_def by auto + thus ?thesis using assms unfolding ordLeq_def ordIso_def + using embed_comp_iso by blast +qed + + +lemma ordIso_ordLeq_trans[trans]: +assumes "r =o r'" and " r' \o r''" +shows "r \o r''" +proof- + have "Well_order r \ Well_order r''" + using assms unfolding ordLeq_def ordIso_def by auto + thus ?thesis using assms unfolding ordLeq_def ordIso_def + using iso_comp_embed by blast +qed + + +lemma ordLess_ordIso_trans[trans]: +assumes "r Well_order r''" + using assms unfolding ordLess_def ordIso_def by auto + thus ?thesis using assms unfolding ordLess_def ordIso_def + using embedS_comp_iso by blast +qed + + +lemma ordIso_ordLess_trans[trans]: +assumes "r =o r'" and " r' Well_order r''" + using assms unfolding ordLess_def ordIso_def by auto + thus ?thesis using assms unfolding ordLess_def ordIso_def + using iso_comp_embedS by blast +qed + + +lemma ordLess_not_embed: +assumes "r (\f'. embed r' r f')" +proof- + obtain f where 1: "Well_order r \ Well_order r'" and 2: "embed r r' f" and + 3: " \ bij_betw f (Field r) (Field r')" + using assms unfolding ordLess_def by (auto simp add: embedS_def) + {fix f' assume *: "embed r' r f'" + hence "bij_betw f (Field r) (Field r')" using 1 2 + by (simp add: embed_bothWays_Field_bij_betw) + with 3 have False by contradiction + } + thus ?thesis by blast +qed + + +lemma ordLess_Field: +assumes OL: "r1 (f`(Field r1) = Field r2)" +proof- + let ?A1 = "Field r1" let ?A2 = "Field r2" + obtain g where + 0: "Well_order r1 \ Well_order r2" and + 1: "embed r1 r2 g \ \(bij_betw g ?A1 ?A2)" + using OL unfolding ordLess_def by (auto simp add: embedS_def) + hence "\a \ ?A1. f a = g a" + using 0 EMB embed_unique[of r1] by auto + hence "\(bij_betw f ?A1 ?A2)" + using 1 bij_betw_cong[of ?A1] by blast + moreover + have "inj_on f ?A1" using EMB 0 by (simp add: embed_inj_on) + ultimately show ?thesis by (simp add: bij_betw_def) +qed + + +lemma ordLess_iff: +"r Well_order r' \ \(\f'. embed r' r f'))" +proof + assume *: "r (\f'. embed r' r f')" using ordLess_not_embed[of r r'] by simp + with * show "Well_order r \ Well_order r' \ \ (\f'. embed r' r f')" + unfolding ordLess_def by auto +next + assume *: "Well_order r \ Well_order r' \ \ (\f'. embed r' r f')" + then obtain f where 1: "embed r r' f" + using wellorders_totally_ordered[of r r'] by blast + moreover + {assume "bij_betw f (Field r) (Field r')" + with * 1 have "embed r' r (inv_into (Field r) f) " + using inv_into_Field_embed_bij_betw[of r r' f] by auto + with * have False by blast + } + ultimately show "(r,r') \ ordLess" + unfolding ordLess_def using * by (fastforce simp add: embedS_def) +qed + + +lemma ordLess_irreflexive: "\ r \(\f. embed r r f)" + unfolding ordLess_iff .. + moreover have "embed r r id" using id_embed[of r] . + ultimately show False by blast +qed + + +lemma ordLeq_iff_ordLess_or_ordIso: +"r \o r' = (r r =o r')" +unfolding ordRels_def embedS_defs iso_defs by blast + + +lemma ordIso_iff_ordLeq: +"(r =o r') = (r \o r' \ r' \o r)" +proof + assume "r =o r'" + then obtain f where 1: "Well_order r \ Well_order r' \ + embed r r' f \ bij_betw f (Field r) (Field r')" + unfolding ordIso_def iso_defs by auto + hence "embed r r' f \ embed r' r (inv_into (Field r) f)" + by (simp add: inv_into_Field_embed_bij_betw) + thus "r \o r' \ r' \o r" + unfolding ordLeq_def using 1 by auto +next + assume "r \o r' \ r' \o r" + then obtain f and g where 1: "Well_order r \ Well_order r' \ + embed r r' f \ embed r' r g" + unfolding ordLeq_def by auto + hence "iso r r' f" by (auto simp add: embed_bothWays_iso) + thus "r =o r'" unfolding ordIso_def using 1 by auto +qed + + +lemma not_ordLess_ordLeq: +"r \ r' \o r" +using ordLess_ordLeq_trans ordLess_irreflexive by blast + + +lemma ordLess_or_ordLeq: +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "r r' \o r" +proof- + have "r \o r' \ r' \o r" + using assms by (simp add: ordLeq_total) + moreover + {assume "\ r r \o r'" + hence "r =o r'" using ordLeq_iff_ordLess_or_ordIso by blast + hence "r' \o r" using ordIso_symmetric ordIso_iff_ordLeq by blast + } + ultimately show ?thesis by blast +qed + + +lemma not_ordLess_ordIso: +"r \ r =o r'" +using assms ordLess_ordIso_trans ordIso_symmetric ordLess_irreflexive by blast + + +lemma not_ordLeq_iff_ordLess: +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "(\ r' \o r) = (r r' o r')" +using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast + + +lemma ordLess_transitive[trans]: +"\r \ r r \o r'" +using ordIso_iff_ordLeq by blast + + +lemma ordLess_imp_ordLeq: +"r r \o r'" +using ordLeq_iff_ordLess_or_ordIso by blast + + +lemma ofilter_subset_ordLeq: +assumes WELL: "Well_order r" and + OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" +shows "(A \ B) = (Restr r A \o Restr r B)" +proof + assume "A \ B" + thus "Restr r A \o Restr r B" + unfolding ordLeq_def using assms + Well_order_Restr Well_order_Restr ofilter_subset_embed by blast +next + assume *: "Restr r A \o Restr r B" + then obtain f where "embed (Restr r A) (Restr r B) f" + unfolding ordLeq_def by blast + {assume "B < A" + hence "Restr r B B" using OFA OFB WELL + wo_rel_def[of r] wo_rel.ofilter_linord[of r A B] by blast +qed + + +lemma ofilter_subset_ordLess: +assumes WELL: "Well_order r" and + OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" +shows "(A < B) = (Restr r A Well_order ?rB" + using WELL Well_order_Restr by blast + have "(A < B) = (\ B \ A)" using assms + wo_rel_def wo_rel.ofilter_linord[of r A B] by blast + also have "\ = (\ Restr r B \o Restr r A)" + using assms ofilter_subset_ordLeq by blast + also have "\ = (Restr r A Well_order r; wo_rel.ofilter r A\ \ (A < Field r) = (Restr r A {}" +shows "Restr r (rel.underS r a) (ofilterIncl r3)" +proof- + have OL13: "r1 Well_order r2 \ Well_order r3" and + 1: "embed r1 r2 f12 \ \(bij_betw f12 ?A1 ?A2)" and + 2: "embed r2 r3 g23 \ \(bij_betw g23 ?A2 ?A3)" + using OL12 OL23 by (auto simp add: ordLess_def embedS_def) + hence "\a \ ?A2. f23 a = g23 a" + using EMB23 embed_unique[of r2 r3] by blast + hence 3: "\(bij_betw f23 ?A2 ?A3)" + using 2 bij_betw_cong[of ?A2 f23 g23] by blast + (* *) + have 4: "wo_rel.ofilter r2 (f12 ` ?A1) \ f12 ` ?A1 \ ?A2" + using 0 1 OL12 by (simp add: embed_Field_ofilter ordLess_Field) + have 5: "wo_rel.ofilter r3 (f23 ` ?A2) \ f23 ` ?A2 \ ?A3" + using 0 EMB23 OL23 by (simp add: embed_Field_ofilter ordLess_Field) + have 6: "wo_rel.ofilter r3 (f13 ` ?A1) \ f13 ` ?A1 \ ?A3" + using 0 EMB13 OL13 by (simp add: embed_Field_ofilter ordLess_Field) + (* *) + have "f12 ` ?A1 < ?A2" + using 0 4 by (auto simp add: wo_rel_def wo_rel.ofilter_def) + moreover have "inj_on f23 ?A2" + using EMB23 0 by (simp add: wo_rel_def embed_inj_on) + ultimately + have "f23 ` (f12 ` ?A1) < f23 ` ?A2" by (simp add: inj_on_strict_subset) + moreover + {have "embed r1 r3 (f23 o f12)" + using 1 EMB23 0 by (auto simp add: comp_embed) + hence "\a \ ?A1. f23(f12 a) = f13 a" + using EMB13 0 embed_unique[of r1 r3 "f23 o f12" f13] by auto + hence "f23 ` (f12 ` ?A1) = f13 ` ?A1" by force + } + ultimately + have "f13 ` ?A1 < f23 ` ?A2" by simp + (* *) + with 5 6 show ?thesis + unfolding ofilterIncl_def by auto +qed + + +lemma ordLess_iff_ordIso_Restr: +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "(r' a \ Field r. r' =o Restr r (rel.underS r a))" +proof(auto) + fix a assume *: "a \ Field r" and **: "r' =o Restr r (rel.underS r a)" + hence "Restr r (rel.underS r a) Well_order r'" and + 2: "embed r' r f \ f ` (Field r') \ Field r" + unfolding ordLess_def embedS_def[abs_def] bij_betw_def using embed_inj_on by blast + hence "wo_rel.ofilter r (f ` (Field r'))" using embed_Field_ofilter by blast + then obtain a where 3: "a \ Field r" and 4: "rel.underS r a = f ` (Field r')" + using 1 2 by (auto simp add: wo_rel.ofilter_underS_Field wo_rel_def) + have "iso r' (Restr r (f ` (Field r'))) f" + using embed_implies_iso_Restr 2 assms by blast + moreover have "Well_order (Restr r (f ` (Field r')))" + using WELL Well_order_Restr by blast + ultimately have "r' =o Restr r (f ` (Field r'))" + using WELL' unfolding ordIso_def by auto + hence "r' =o Restr r (rel.underS r a)" using 4 by auto + thus "\a \ Field r. r' =o Restr r (rel.underS r a)" using 3 by auto +qed + + +lemma internalize_ordLess: +"(r' p. Field p < Field r \ r' =o p \ p Well_order r'" unfolding ordLess_def by auto + with * obtain a where 1: "a \ Field r" and 2: "r' =o Restr r (rel.underS r a)" + using ordLess_iff_ordIso_Restr by blast + let ?p = "Restr r (rel.underS r a)" + have "wo_rel.ofilter r (rel.underS r a)" using 0 + by (simp add: wo_rel_def wo_rel.underS_ofilter) + hence "Field ?p = rel.underS r a" using 0 Field_Restr_ofilter by blast + hence "Field ?p < Field r" using rel.underS_Field2 1 by fastforce + moreover have "?p p. Field p < Field r \ r' =o p \ p p. Field p < Field r \ r' =o p \ p o r) = (\p. Field p \ Field r \ r' =o p \ p \o r)" +proof + assume *: "r' \o r" + moreover + {assume "r' r' =o p \ p p. Field p \ Field r \ r' =o p \ p \o r" + using ordLeq_iff_ordLess_or_ordIso by blast + } + moreover + have "r \o r" using * ordLeq_def ordLeq_reflexive by blast + ultimately show "\p. Field p \ Field r \ r' =o p \ p \o r" + using ordLeq_iff_ordLess_or_ordIso by blast +next + assume "\p. Field p \ Field r \ r' =o p \ p \o r" + thus "r' \o r" using ordIso_ordLeq_trans by blast +qed + + +lemma ordLeq_iff_ordLess_Restr: +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "(r \o r') = (\a \ Field r. Restr r (rel.underS r a) o r'" + fix a assume "a \ Field r" + hence "Restr r (rel.underS r a) a \ Field r. Restr r (rel.underS r a) Field r \ r' =o Restr r (rel.underS r a)" + using assms ordLess_iff_ordIso_Restr by blast + hence False using * not_ordLess_ordIso ordIso_symmetric by blast + } + thus "r \o r'" using ordLess_or_ordLeq assms by blast +qed + + +lemma finite_ordLess_infinite: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + FIN: "finite(Field r)" and INF: "infinite(Field r')" +shows "r o r" + then obtain h where "inj_on h (Field r') \ h ` (Field r') \ Field r" + unfolding ordLeq_def using assms embed_inj_on embed_Field by blast + hence False using finite_imageD finite_subset FIN INF by metis + } + thus ?thesis using WELL WELL' ordLess_or_ordLeq by blast +qed + + +lemma finite_well_order_on_ordIso: +assumes FIN: "finite A" and + WELL: "well_order_on A r" and WELL': "well_order_on A r'" +shows "r =o r'" +proof- + have 0: "Well_order r \ Well_order r' \ Field r = A \ Field r' = A" + using assms rel.well_order_on_Well_order by blast + moreover + have "\r r'. well_order_on A r \ well_order_on A r' \ r \o r' + \ r =o r'" + proof(clarify) + fix r r' assume *: "well_order_on A r" and **: "well_order_on A r'" + have 2: "Well_order r \ Well_order r' \ Field r = A \ Field r' = A" + using * ** rel.well_order_on_Well_order by blast + assume "r \o r'" + then obtain f where 1: "embed r r' f" and + "inj_on f A \ f ` A \ A" + unfolding ordLeq_def using 2 embed_inj_on embed_Field by blast + hence "bij_betw f A A" unfolding bij_betw_def using FIN endo_inj_surj by blast + thus "r =o r'" unfolding ordIso_def iso_def[abs_def] using 1 2 by auto + qed + ultimately show ?thesis using assms ordLeq_total ordIso_symmetric by blast +qed + + +subsection{* @{text " 'a rel \ 'a set" +where "ord_to_filter r0 r \ (SOME f. embed r r0 f) ` (Field r)" + + +lemma ord_to_filter_compat: +"compat (ordLess Int (ordLess^-1``{r0} \ ordLess^-1``{r0})) + (ofilterIncl r0) + (ord_to_filter r0)" +proof(unfold compat_def ord_to_filter_def, clarify) + fix r1::"'a rel" and r2::"'a rel" + let ?A1 = "Field r1" let ?A2 ="Field r2" let ?A0 ="Field r0" + let ?phi10 = "\ f10. embed r1 r0 f10" let ?f10 = "SOME f. ?phi10 f" + let ?phi20 = "\ f20. embed r2 r0 f20" let ?f20 = "SOME f. ?phi20 f" + assume *: "r1 f. ?phi10 f) \ (\f. ?phi20 f)" + by (auto simp add: ordLess_def embedS_def) + hence "?phi10 ?f10 \ ?phi20 ?f20" by (auto simp add: someI_ex) + thus "(?f10 ` ?A1, ?f20 ` ?A2) \ ofilterIncl r0" + using * ** by (simp add: embed_ordLess_ofilterIncl) +qed + + +theorem wf_ordLess: "wf ordLess" +proof- + {fix r0 :: "('a \ 'a) set" + (* need to annotate here!*) + let ?ordLess = "ordLess::('d rel * 'd rel) set" + let ?R = "?ordLess Int (?ordLess^-1``{r0} \ ?ordLess^-1``{r0})" + {assume Case1: "Well_order r0" + hence "wf ?R" + using wf_ofilterIncl[of r0] + compat_wf[of ?R "ofilterIncl r0" "ord_to_filter r0"] + ord_to_filter_compat[of r0] by auto + } + moreover + {assume Case2: "\ Well_order r0" + hence "?R = {}" unfolding ordLess_def by auto + hence "wf ?R" using wf_empty by simp + } + ultimately have "wf ?R" by blast + } + thus ?thesis by (simp add: trans_wf_iff ordLess_trans) +qed + +corollary exists_minim_Well_order: +assumes NE: "R \ {}" and WELL: "\r \ R. Well_order r" +shows "\r \ R. \r' \ R. r \o r'" +proof- + obtain r where "r \ R \ (\r' \ R. \ r' ('a \ 'a') \ 'a' rel" +where +"dir_image r f = {(f a, f b)| a b. (a,b) \ r}" + + +lemma dir_image_Field: +"Field(dir_image r f) \ f ` (Field r)" +unfolding dir_image_def Field_def by auto + + +lemma dir_image_minus_Id: +"inj_on f (Field r) \ (dir_image r f) - Id = dir_image (r - Id) f" +unfolding inj_on_def Field_def dir_image_def by auto + + +lemma Refl_dir_image: +assumes "Refl r" +shows "Refl(dir_image r f)" +proof- + {fix a' b' + assume "(a',b') \ dir_image r f" + then obtain a b where 1: "a' = f a \ b' = f b \ (a,b) \ r" + unfolding dir_image_def by blast + hence "a \ Field r \ b \ Field r" using Field_def by fastforce + hence "(a,a) \ r \ (b,b) \ r" using assms by (simp add: refl_on_def) + with 1 have "(a',a') \ dir_image r f \ (b',b') \ dir_image r f" + unfolding dir_image_def by auto + } + thus ?thesis + by(unfold refl_on_def Field_def Domain_def Range_def, auto) +qed + + +lemma trans_dir_image: +assumes TRANS: "trans r" and INJ: "inj_on f (Field r)" +shows "trans(dir_image r f)" +proof(unfold trans_def, auto) + fix a' b' c' + assume "(a',b') \ dir_image r f" "(b',c') \ dir_image r f" + then obtain a b1 b2 c where 1: "a' = f a \ b' = f b1 \ b' = f b2 \ c' = f c" and + 2: "(a,b1) \ r \ (b2,c) \ r" + unfolding dir_image_def by blast + hence "b1 \ Field r \ b2 \ Field r" + unfolding Field_def by auto + hence "b1 = b2" using 1 INJ unfolding inj_on_def by auto + hence "(a,c): r" using 2 TRANS unfolding trans_def by blast + thus "(a',c') \ dir_image r f" + unfolding dir_image_def using 1 by auto +qed + + +lemma Preorder_dir_image: +"\Preorder r; inj_on f (Field r)\ \ Preorder (dir_image r f)" +by (simp add: preorder_on_def Refl_dir_image trans_dir_image) + + +lemma antisym_dir_image: +assumes AN: "antisym r" and INJ: "inj_on f (Field r)" +shows "antisym(dir_image r f)" +proof(unfold antisym_def, auto) + fix a' b' + assume "(a',b') \ dir_image r f" "(b',a') \ dir_image r f" + then obtain a1 b1 a2 b2 where 1: "a' = f a1 \ a' = f a2 \ b' = f b1 \ b' = f b2" and + 2: "(a1,b1) \ r \ (b2,a2) \ r " and + 3: "{a1,a2,b1,b2} \ Field r" + unfolding dir_image_def Field_def by blast + hence "a1 = a2 \ b1 = b2" using INJ unfolding inj_on_def by auto + hence "a1 = b2" using 2 AN unfolding antisym_def by auto + thus "a' = b'" using 1 by auto +qed + + +lemma Partial_order_dir_image: +"\Partial_order r; inj_on f (Field r)\ \ Partial_order (dir_image r f)" +by (simp add: partial_order_on_def Preorder_dir_image antisym_dir_image) + + +lemma Total_dir_image: +assumes TOT: "Total r" and INJ: "inj_on f (Field r)" +shows "Total(dir_image r f)" +proof(unfold total_on_def, intro ballI impI) + fix a' b' + assume "a' \ Field (dir_image r f)" "b' \ Field (dir_image r f)" + then obtain a and b where 1: "a \ Field r \ b \ Field r \ f a = a' \ f b = b'" + using dir_image_Field[of r f] by blast + moreover assume "a' \ b'" + ultimately have "a \ b" using INJ unfolding inj_on_def by auto + hence "(a,b) \ r \ (b,a) \ r" using 1 TOT unfolding total_on_def by auto + thus "(a',b') \ dir_image r f \ (b',a') \ dir_image r f" + using 1 unfolding dir_image_def by auto +qed + + +lemma Linear_order_dir_image: +"\Linear_order r; inj_on f (Field r)\ \ Linear_order (dir_image r f)" +by (simp add: linear_order_on_def Partial_order_dir_image Total_dir_image) + + +lemma wf_dir_image: +assumes WF: "wf r" and INJ: "inj_on f (Field r)" +shows "wf(dir_image r f)" +proof(unfold wf_eq_minimal2, intro allI impI, elim conjE) + fix A'::"'b set" + assume SUB: "A' \ Field(dir_image r f)" and NE: "A' \ {}" + obtain A where A_def: "A = {a \ Field r. f a \ A'}" by blast + have "A \ {} \ A \ Field r" + using A_def dir_image_Field[of r f] SUB NE by blast + then obtain a where 1: "a \ A \ (\b \ A. (b,a) \ r)" + using WF unfolding wf_eq_minimal2 by metis + have "\b' \ A'. (b',f a) \ dir_image r f" + proof(clarify) + fix b' assume *: "b' \ A'" and **: "(b',f a) \ dir_image r f" + obtain b1 a1 where 2: "b' = f b1 \ f a = f a1" and + 3: "(b1,a1) \ r \ {a1,b1} \ Field r" + using ** unfolding dir_image_def Field_def by blast + hence "a = a1" using 1 A_def INJ unfolding inj_on_def by auto + hence "b1 \ A \ (b1,a) \ r" using 2 3 A_def * by auto + with 1 show False by auto + qed + thus "\a'\A'. \b'\A'. (b', a') \ dir_image r f" + using A_def 1 by blast +qed + + +lemma Well_order_dir_image: +"\Well_order r; inj_on f (Field r)\ \ Well_order (dir_image r f)" +using assms unfolding well_order_on_def +using Linear_order_dir_image[of r f] wf_dir_image[of "r - Id" f] + dir_image_minus_Id[of f r] + subset_inj_on[of f "Field r" "Field(r - Id)"] + mono_Field[of "r - Id" r] by auto + + +lemma dir_image_Field2: +"Refl r \ Field(dir_image r f) = f ` (Field r)" +unfolding Field_def dir_image_def refl_on_def Domain_def Range_def by blast + + +lemma dir_image_bij_betw: +"\Well_order r; inj_on f (Field r)\ \ bij_betw f (Field r) (Field (dir_image r f))" +unfolding bij_betw_def +by (simp add: dir_image_Field2 order_on_defs) + + +lemma dir_image_compat: +"compat r (dir_image r f) f" +unfolding compat_def dir_image_def by auto + + +lemma dir_image_iso: +"\Well_order r; inj_on f (Field r)\ \ iso r (dir_image r f) f" +using iso_iff3 dir_image_compat dir_image_bij_betw Well_order_dir_image by blast + + +lemma dir_image_ordIso: +"\Well_order r; inj_on f (Field r)\ \ r =o dir_image r f" +unfolding ordIso_def using dir_image_iso Well_order_dir_image by blast + + +lemma Well_order_iso_copy: +assumes WELL: "well_order_on A r" and BIJ: "bij_betw f A A'" +shows "\r'. well_order_on A' r' \ r =o r'" +proof- + let ?r' = "dir_image r f" + have 1: "A = Field r \ Well_order r" + using WELL rel.well_order_on_Well_order by blast + hence 2: "iso r ?r' f" + using dir_image_iso using BIJ unfolding bij_betw_def by auto + hence "f ` (Field r) = Field ?r'" using 1 iso_iff[of r ?r'] by blast + hence "Field ?r' = A'" + using 1 BIJ unfolding bij_betw_def by auto + moreover have "Well_order ?r'" + using 1 Well_order_dir_image BIJ unfolding bij_betw_def by blast + ultimately show ?thesis unfolding ordIso_def using 1 2 by blast +qed + + + +subsection {* Bounded square *} + + +text{* This construction essentially defines, for an order relation @{text "r"}, a lexicographic +order @{text "bsqr r"} on @{text "(Field r) \ (Field r)"}, applying the +following criteria (in this order): +\begin{itemize} +\item compare the maximums; +\item compare the first components; +\item compare the second components. +\end{itemize} +% +The only application of this construction that we are aware of is +at proving that the square of an infinite set has the same cardinal +as that set. The essential property required there (and which is ensured by this +construction) is that any proper order filter of the product order is included in a rectangle, i.e., +in a product of proper filters on the original relation (assumed to be a well-order). *} + + +definition bsqr :: "'a rel => ('a * 'a)rel" +where +"bsqr r = {((a1,a2),(b1,b2)). + {a1,a2,b1,b2} \ Field r \ + (a1 = b1 \ a2 = b2 \ + (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ + wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ + wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id + )}" + + +lemma Field_bsqr: +"Field (bsqr r) = Field r \ Field r" +proof + show "Field (bsqr r) \ Field r \ Field r" + proof- + {fix a1 a2 assume "(a1,a2) \ Field (bsqr r)" + moreover + have "\ b1 b2. ((a1,a2),(b1,b2)) \ bsqr r \ ((b1,b2),(a1,a2)) \ bsqr r \ + a1 \ Field r \ a2 \ Field r" unfolding bsqr_def by auto + ultimately have "a1 \ Field r \ a2 \ Field r" unfolding Field_def by auto + } + thus ?thesis unfolding Field_def by force + qed +next + show "Field r \ Field r \ Field (bsqr r)" + proof(auto) + fix a1 a2 assume "a1 \ Field r" and "a2 \ Field r" + hence "((a1,a2),(a1,a2)) \ bsqr r" unfolding bsqr_def by blast + thus "(a1,a2) \ Field (bsqr r)" unfolding Field_def by auto + qed +qed + + +lemma bsqr_Refl: "Refl(bsqr r)" +by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def) + + +lemma bsqr_Trans: +assumes "Well_order r" +shows "trans (bsqr r)" +proof(unfold trans_def, auto) + (* Preliminary facts *) + have Well: "wo_rel r" using assms wo_rel_def by auto + hence Trans: "trans r" using wo_rel.TRANS by auto + have Anti: "antisym r" using wo_rel.ANTISYM Well by auto + hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id) + (* Main proof *) + fix a1 a2 b1 b2 c1 c2 + assume *: "((a1,a2),(b1,b2)) \ bsqr r" and **: "((b1,b2),(c1,c2)) \ bsqr r" + hence 0: "{a1,a2,b1,b2,c1,c2} \ Field r" unfolding bsqr_def by auto + have 1: "a1 = b1 \ a2 = b2 \ (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ + wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ + wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" + using * unfolding bsqr_def by auto + have 2: "b1 = c1 \ b2 = c2 \ (wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id \ + wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id \ + wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1 \ (b2,c2) \ r - Id" + using ** unfolding bsqr_def by auto + show "((a1,a2),(c1,c2)) \ bsqr r" + proof- + {assume Case1: "a1 = b1 \ a2 = b2" + hence ?thesis using ** by simp + } + moreover + {assume Case2: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id" + {assume Case21: "b1 = c1 \ b2 = c2" + hence ?thesis using * by simp + } + moreover + {assume Case22: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" + hence "(wo_rel.max2 r a1 a2, wo_rel.max2 r c1 c2) \ r - Id" + using Case2 TransS trans_def[of "r - Id"] by blast + hence ?thesis using 0 unfolding bsqr_def by auto + } + moreover + {assume Case23_4: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2" + hence ?thesis using Case2 0 unfolding bsqr_def by auto + } + ultimately have ?thesis using 0 2 by auto + } + moreover + {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id" + {assume Case31: "b1 = c1 \ b2 = c2" + hence ?thesis using * by simp + } + moreover + {assume Case32: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" + hence ?thesis using Case3 0 unfolding bsqr_def by auto + } + moreover + {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id" + hence "(a1,c1) \ r - Id" + using Case3 TransS trans_def[of "r - Id"] by blast + hence ?thesis using Case3 Case33 0 unfolding bsqr_def by auto + } + moreover + {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1" + hence ?thesis using Case3 0 unfolding bsqr_def by auto + } + ultimately have ?thesis using 0 2 by auto + } + moreover + {assume Case4: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" + {assume Case41: "b1 = c1 \ b2 = c2" + hence ?thesis using * by simp + } + moreover + {assume Case42: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" + hence ?thesis using Case4 0 unfolding bsqr_def by force + } + moreover + {assume Case43: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id" + hence ?thesis using Case4 0 unfolding bsqr_def by auto + } + moreover + {assume Case44: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1 \ (b2,c2) \ r - Id" + hence "(a2,c2) \ r - Id" + using Case4 TransS trans_def[of "r - Id"] by blast + hence ?thesis using Case4 Case44 0 unfolding bsqr_def by auto + } + ultimately have ?thesis using 0 2 by auto + } + ultimately show ?thesis using 0 1 by auto + qed +qed + + +lemma bsqr_antisym: +assumes "Well_order r" +shows "antisym (bsqr r)" +proof(unfold antisym_def, clarify) + (* Preliminary facts *) + have Well: "wo_rel r" using assms wo_rel_def by auto + hence Trans: "trans r" using wo_rel.TRANS by auto + have Anti: "antisym r" using wo_rel.ANTISYM Well by auto + hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id) + hence IrrS: "\a b. \((a,b) \ r - Id \ (b,a) \ r - Id)" + using Anti trans_def[of "r - Id"] antisym_def[of "r - Id"] by blast + (* Main proof *) + fix a1 a2 b1 b2 + assume *: "((a1,a2),(b1,b2)) \ bsqr r" and **: "((b1,b2),(a1,a2)) \ bsqr r" + hence 0: "{a1,a2,b1,b2} \ Field r" unfolding bsqr_def by auto + have 1: "a1 = b1 \ a2 = b2 \ (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ + wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ + wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" + using * unfolding bsqr_def by auto + have 2: "b1 = a1 \ b2 = a2 \ (wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id \ + wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \ (b1,a1) \ r - Id \ + wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \ b1 = a1 \ (b2,a2) \ r - Id" + using ** unfolding bsqr_def by auto + show "a1 = b1 \ a2 = b2" + proof- + {assume Case1: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id" + {assume Case11: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" + hence False using Case1 IrrS by blast + } + moreover + {assume Case12_3: "wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2" + hence False using Case1 by auto + } + ultimately have ?thesis using 0 2 by auto + } + moreover + {assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id" + {assume Case21: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" + hence False using Case2 by auto + } + moreover + {assume Case22: "(b1,a1) \ r - Id" + hence False using Case2 IrrS by blast + } + moreover + {assume Case23: "b1 = a1" + hence False using Case2 by auto + } + ultimately have ?thesis using 0 2 by auto + } + moreover + {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" + moreover + {assume Case31: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" + hence False using Case3 by auto + } + moreover + {assume Case32: "(b1,a1) \ r - Id" + hence False using Case3 by auto + } + moreover + {assume Case33: "(b2,a2) \ r - Id" + hence False using Case3 IrrS by blast + } + ultimately have ?thesis using 0 2 by auto + } + ultimately show ?thesis using 0 1 by blast + qed +qed + + +lemma bsqr_Total: +assumes "Well_order r" +shows "Total(bsqr r)" +proof- + (* Preliminary facts *) + have Well: "wo_rel r" using assms wo_rel_def by auto + hence Total: "\a \ Field r. \b \ Field r. (a,b) \ r \ (b,a) \ r" + using wo_rel.TOTALS by auto + (* Main proof *) + {fix a1 a2 b1 b2 assume "{(a1,a2), (b1,b2)} \ Field(bsqr r)" + hence 0: "a1 \ Field r \ a2 \ Field r \ b1 \ Field r \ b2 \ Field r" + using Field_bsqr by blast + have "((a1,a2) = (b1,b2) \ ((a1,a2),(b1,b2)) \ bsqr r \ ((b1,b2),(a1,a2)) \ bsqr r)" + proof(rule wo_rel.cases_Total[of r a1 a2], clarsimp simp add: Well, simp add: 0) + (* Why didn't clarsimp simp add: Well 0 do the same job? *) + assume Case1: "(a1,a2) \ r" + hence 1: "wo_rel.max2 r a1 a2 = a2" + using Well 0 by (simp add: wo_rel.max2_equals2) + show ?thesis + proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0) + assume Case11: "(b1,b2) \ r" + hence 2: "wo_rel.max2 r b1 b2 = b2" + using Well 0 by (simp add: wo_rel.max2_equals2) + show ?thesis + proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) + assume Case111: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" + thus ?thesis using 0 1 2 unfolding bsqr_def by auto + next + assume Case112: "a2 = b2" + show ?thesis + proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) + assume Case1121: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" + thus ?thesis using 0 1 2 Case112 unfolding bsqr_def by auto + next + assume Case1122: "a1 = b1" + thus ?thesis using Case112 by auto + qed + qed + next + assume Case12: "(b2,b1) \ r" + hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1) + show ?thesis + proof(rule wo_rel.cases_Total3[of r a2 b1], clarsimp simp add: Well, simp add: 0) + assume Case121: "(a2,b1) \ r - Id \ (b1,a2) \ r - Id" + thus ?thesis using 0 1 3 unfolding bsqr_def by auto + next + assume Case122: "a2 = b1" + show ?thesis + proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) + assume Case1221: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" + thus ?thesis using 0 1 3 Case122 unfolding bsqr_def by auto + next + assume Case1222: "a1 = b1" + show ?thesis + proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) + assume Case12221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" + thus ?thesis using 0 1 3 Case122 Case1222 unfolding bsqr_def by auto + next + assume Case12222: "a2 = b2" + thus ?thesis using Case122 Case1222 by auto + qed + qed + qed + qed + next + assume Case2: "(a2,a1) \ r" + hence 1: "wo_rel.max2 r a1 a2 = a1" using Well 0 by (simp add: wo_rel.max2_equals1) + show ?thesis + proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0) + assume Case21: "(b1,b2) \ r" + hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2) + show ?thesis + proof(rule wo_rel.cases_Total3[of r a1 b2], clarsimp simp add: Well, simp add: 0) + assume Case211: "(a1,b2) \ r - Id \ (b2,a1) \ r - Id" + thus ?thesis using 0 1 2 unfolding bsqr_def by auto + next + assume Case212: "a1 = b2" + show ?thesis + proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) + assume Case2121: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" + thus ?thesis using 0 1 2 Case212 unfolding bsqr_def by auto + next + assume Case2122: "a1 = b1" + show ?thesis + proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) + assume Case21221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" + thus ?thesis using 0 1 2 Case212 Case2122 unfolding bsqr_def by auto + next + assume Case21222: "a2 = b2" + thus ?thesis using Case2122 Case212 by auto + qed + qed + qed + next + assume Case22: "(b2,b1) \ r" + hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1) + show ?thesis + proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) + assume Case221: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" + thus ?thesis using 0 1 3 unfolding bsqr_def by auto + next + assume Case222: "a1 = b1" + show ?thesis + proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) + assume Case2221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" + thus ?thesis using 0 1 3 Case222 unfolding bsqr_def by auto + next + assume Case2222: "a2 = b2" + thus ?thesis using Case222 by auto + qed + qed + qed + qed + } + thus ?thesis unfolding total_on_def by fast +qed + + +lemma bsqr_Linear_order: +assumes "Well_order r" +shows "Linear_order(bsqr r)" +unfolding order_on_defs +using assms bsqr_Refl bsqr_Trans bsqr_antisym bsqr_Total by blast + + +lemma bsqr_Well_order: +assumes "Well_order r" +shows "Well_order(bsqr r)" +using assms +proof(simp add: bsqr_Linear_order Linear_order_Well_order_iff, intro allI impI) + have 0: "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r)" + using assms well_order_on_def Linear_order_Well_order_iff by blast + fix D assume *: "D \ Field (bsqr r)" and **: "D \ {}" + hence 1: "D \ Field r \ Field r" unfolding Field_bsqr by simp + (* *) + obtain M where M_def: "M = {wo_rel.max2 r a1 a2| a1 a2. (a1,a2) \ D}" by blast + have "M \ {}" using 1 M_def ** by auto + moreover + have "M \ Field r" unfolding M_def + using 1 assms wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce + ultimately obtain m where m_min: "m \ M \ (\a \ M. (m,a) \ r)" + using 0 by blast + (* *) + obtain A1 where A1_def: "A1 = {a1. \a2. (a1,a2) \ D \ wo_rel.max2 r a1 a2 = m}" by blast + have "A1 \ Field r" unfolding A1_def using 1 by auto + moreover have "A1 \ {}" unfolding A1_def using m_min unfolding M_def by blast + ultimately obtain a1 where a1_min: "a1 \ A1 \ (\a \ A1. (a1,a) \ r)" + using 0 by blast + (* *) + obtain A2 where A2_def: "A2 = {a2. (a1,a2) \ D \ wo_rel.max2 r a1 a2 = m}" by blast + have "A2 \ Field r" unfolding A2_def using 1 by auto + moreover have "A2 \ {}" unfolding A2_def + using m_min a1_min unfolding A1_def M_def by blast + ultimately obtain a2 where a2_min: "a2 \ A2 \ (\a \ A2. (a2,a) \ r)" + using 0 by blast + (* *) + have 2: "wo_rel.max2 r a1 a2 = m" + using a1_min a2_min unfolding A1_def A2_def by auto + have 3: "(a1,a2) \ D" using a2_min unfolding A2_def by auto + (* *) + moreover + {fix b1 b2 assume ***: "(b1,b2) \ D" + hence 4: "{a1,a2,b1,b2} \ Field r" using 1 3 by blast + have 5: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r" + using *** a1_min a2_min m_min unfolding A1_def A2_def M_def by auto + have "((a1,a2),(b1,b2)) \ bsqr r" + proof(cases "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2") + assume Case1: "wo_rel.max2 r a1 a2 \ wo_rel.max2 r b1 b2" + thus ?thesis unfolding bsqr_def using 4 5 by auto + next + assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2" + hence "b1 \ A1" unfolding A1_def using 2 *** by auto + hence 6: "(a1,b1) \ r" using a1_min by auto + show ?thesis + proof(cases "a1 = b1") + assume Case21: "a1 \ b1" + thus ?thesis unfolding bsqr_def using 4 Case2 6 by auto + next + assume Case22: "a1 = b1" + hence "b2 \ A2" unfolding A2_def using 2 *** Case2 by auto + hence 7: "(a2,b2) \ r" using a2_min by auto + thus ?thesis unfolding bsqr_def using 4 7 Case2 Case22 by auto + qed + qed + } + (* *) + ultimately show "\d \ D. \d' \ D. (d,d') \ bsqr r" by fastforce +qed + + +lemma bsqr_max2: +assumes WELL: "Well_order r" and LEQ: "((a1,a2),(b1,b2)) \ bsqr r" +shows "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r" +proof- + have "{(a1,a2),(b1,b2)} \ Field(bsqr r)" + using LEQ unfolding Field_def by auto + hence "{a1,a2,b1,b2} \ Field r" unfolding Field_bsqr by auto + hence "{wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2} \ Field r" + using WELL wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce + moreover have "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2" + using LEQ unfolding bsqr_def by auto + ultimately show ?thesis using WELL unfolding order_on_defs refl_on_def by auto +qed + + +lemma bsqr_ofilter: +assumes WELL: "Well_order r" and + OF: "wo_rel.ofilter (bsqr r) D" and SUB: "D < Field r \ Field r" and + NE: "\ (\a. Field r = rel.under r a)" +shows "\A. wo_rel.ofilter r A \ A < Field r \ D \ A \ A" +proof- + let ?r' = "bsqr r" + have Well: "wo_rel r" using WELL wo_rel_def by blast + hence Trans: "trans r" using wo_rel.TRANS by blast + have Well': "Well_order ?r' \ wo_rel ?r'" + using WELL bsqr_Well_order wo_rel_def by blast + (* *) + have "D < Field ?r'" unfolding Field_bsqr using SUB . + with OF obtain a1 and a2 where + "(a1,a2) \ Field ?r'" and 1: "D = rel.underS ?r' (a1,a2)" + using Well' wo_rel.ofilter_underS_Field[of ?r' D] by auto + hence 2: "{a1,a2} \ Field r" unfolding Field_bsqr by auto + let ?m = "wo_rel.max2 r a1 a2" + have "D \ (rel.under r ?m) \ (rel.under r ?m)" + proof(unfold 1) + {fix b1 b2 + let ?n = "wo_rel.max2 r b1 b2" + assume "(b1,b2) \ rel.underS ?r' (a1,a2)" + hence 3: "((b1,b2),(a1,a2)) \ ?r'" + unfolding rel.underS_def by blast + hence "(?n,?m) \ r" using WELL by (simp add: bsqr_max2) + moreover + {have "(b1,b2) \ Field ?r'" using 3 unfolding Field_def by auto + hence "{b1,b2} \ Field r" unfolding Field_bsqr by auto + hence "(b1,?n) \ r \ (b2,?n) \ r" + using Well by (simp add: wo_rel.max2_greater) + } + ultimately have "(b1,?m) \ r \ (b2,?m) \ r" + using Trans trans_def[of r] by blast + hence "(b1,b2) \ (rel.under r ?m) \ (rel.under r ?m)" unfolding rel.under_def by simp} + thus "rel.underS ?r' (a1,a2) \ (rel.under r ?m) \ (rel.under r ?m)" by auto + qed + moreover have "wo_rel.ofilter r (rel.under r ?m)" + using Well by (simp add: wo_rel.under_ofilter) + moreover have "rel.under r ?m < Field r" + using NE rel.under_Field[of r ?m] by blast + ultimately show ?thesis by blast +qed + + +end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Constructions_on_Wellorders_LFP.thy --- a/src/HOL/Cardinals/Constructions_on_Wellorders_LFP.thy Mon Nov 18 18:04:45 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1621 +0,0 @@ -(* Title: HOL/Cardinals/Constructions_on_Wellorders_LFP.thy - Author: Andrei Popescu, TU Muenchen - Copyright 2012 - -Constructions on wellorders (LFP). -*) - -header {* Constructions on Wellorders (LFP) *} - -theory Constructions_on_Wellorders_LFP -imports Wellorder_Embedding_LFP -begin - - -text {* In this section, we study basic constructions on well-orders, such as restriction to -a set/order filter, copy via direct images, ordinal-like sum of disjoint well-orders, -and bounded square. We also define between well-orders -the relations @{text "ordLeq"}, of being embedded (abbreviated @{text "\o"}), -@{text "ordLess"}, of being strictly embedded (abbreviated @{text " 'a set \ 'a rel" -where "Restr r A \ r Int (A \ A)" - - -lemma Restr_subset: -"A \ B \ Restr (Restr r B) A = Restr r A" -by blast - - -lemma Restr_Field: "Restr r (Field r) = r" -unfolding Field_def by auto - - -lemma Refl_Restr: "Refl r \ Refl(Restr r A)" -unfolding refl_on_def Field_def by auto - - -lemma antisym_Restr: -"antisym r \ antisym(Restr r A)" -unfolding antisym_def Field_def by auto - - -lemma Total_Restr: -"Total r \ Total(Restr r A)" -unfolding total_on_def Field_def by auto - - -lemma trans_Restr: -"trans r \ trans(Restr r A)" -unfolding trans_def Field_def by blast - - -lemma Preorder_Restr: -"Preorder r \ Preorder(Restr r A)" -unfolding preorder_on_def by (simp add: Refl_Restr trans_Restr) - - -lemma Partial_order_Restr: -"Partial_order r \ Partial_order(Restr r A)" -unfolding partial_order_on_def by (simp add: Preorder_Restr antisym_Restr) - - -lemma Linear_order_Restr: -"Linear_order r \ Linear_order(Restr r A)" -unfolding linear_order_on_def by (simp add: Partial_order_Restr Total_Restr) - - -lemma Well_order_Restr: -assumes "Well_order r" -shows "Well_order(Restr r A)" -proof- - have "Restr r A - Id \ r - Id" using Restr_subset by blast - hence "wf(Restr r A - Id)" using assms - using well_order_on_def wf_subset by blast - thus ?thesis using assms unfolding well_order_on_def - by (simp add: Linear_order_Restr) -qed - - -lemma Field_Restr_subset: "Field(Restr r A) \ A" -by (auto simp add: Field_def) - - -lemma Refl_Field_Restr: -"Refl r \ Field(Restr r A) = (Field r) Int A" -by (auto simp add: refl_on_def Field_def) - - -lemma Refl_Field_Restr2: -"\Refl r; A \ Field r\ \ Field(Restr r A) = A" -by (auto simp add: Refl_Field_Restr) - - -lemma well_order_on_Restr: -assumes WELL: "Well_order r" and SUB: "A \ Field r" -shows "well_order_on A (Restr r A)" -using assms -using Well_order_Restr[of r A] Refl_Field_Restr2[of r A] - order_on_defs[of "Field r" r] by auto - - -subsection {* Order filters versus restrictions and embeddings *} - - -lemma Field_Restr_ofilter: -"\Well_order r; wo_rel.ofilter r A\ \ Field(Restr r A) = A" -by (auto simp add: wo_rel_def wo_rel.ofilter_def wo_rel.REFL Refl_Field_Restr2) - - -lemma ofilter_Restr_under: -assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and IN: "a \ A" -shows "rel.under (Restr r A) a = rel.under r a" -using assms wo_rel_def -proof(auto simp add: wo_rel.ofilter_def rel.under_def) - fix b assume *: "a \ A" and "(b,a) \ r" - hence "b \ rel.under r a \ a \ Field r" - unfolding rel.under_def using Field_def by fastforce - thus "b \ A" using * assms by (auto simp add: wo_rel_def wo_rel.ofilter_def) -qed - - -lemma ofilter_embed: -assumes "Well_order r" -shows "wo_rel.ofilter r A = (A \ Field r \ embed (Restr r A) r id)" -proof - assume *: "wo_rel.ofilter r A" - show "A \ Field r \ embed (Restr r A) r id" - proof(unfold embed_def, auto) - fix a assume "a \ A" thus "a \ Field r" using assms * - by (auto simp add: wo_rel_def wo_rel.ofilter_def) - next - fix a assume "a \ Field (Restr r A)" - thus "bij_betw id (rel.under (Restr r A) a) (rel.under r a)" using assms * - by (simp add: ofilter_Restr_under Field_Restr_ofilter) - qed -next - assume *: "A \ Field r \ embed (Restr r A) r id" - hence "Field(Restr r A) \ Field r" - using assms embed_Field[of "Restr r A" r id] id_def - Well_order_Restr[of r] by auto - {fix a assume "a \ A" - hence "a \ Field(Restr r A)" using * assms - by (simp add: order_on_defs Refl_Field_Restr2) - hence "bij_betw id (rel.under (Restr r A) a) (rel.under r a)" - using * unfolding embed_def by auto - hence "rel.under r a \ rel.under (Restr r A) a" - unfolding bij_betw_def by auto - also have "\ \ Field(Restr r A)" by (simp add: rel.under_Field) - also have "\ \ A" by (simp add: Field_Restr_subset) - finally have "rel.under r a \ A" . - } - thus "wo_rel.ofilter r A" using assms * by (simp add: wo_rel_def wo_rel.ofilter_def) -qed - - -lemma ofilter_Restr_Int: -assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" -shows "wo_rel.ofilter (Restr r B) (A Int B)" -proof- - let ?rB = "Restr r B" - have Well: "wo_rel r" unfolding wo_rel_def using WELL . - hence Refl: "Refl r" by (simp add: wo_rel.REFL) - hence Field: "Field ?rB = Field r Int B" - using Refl_Field_Restr by blast - have WellB: "wo_rel ?rB \ Well_order ?rB" using WELL - by (simp add: Well_order_Restr wo_rel_def) - (* Main proof *) - show ?thesis using WellB assms - proof(auto simp add: wo_rel.ofilter_def rel.under_def) - fix a assume "a \ A" and *: "a \ B" - hence "a \ Field r" using OFA Well by (auto simp add: wo_rel.ofilter_def) - with * show "a \ Field ?rB" using Field by auto - next - fix a b assume "a \ A" and "(b,a) \ r" - thus "b \ A" using Well OFA by (auto simp add: wo_rel.ofilter_def rel.under_def) - qed -qed - - -lemma ofilter_Restr_subset: -assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and SUB: "A \ B" -shows "wo_rel.ofilter (Restr r B) A" -proof- - have "A Int B = A" using SUB by blast - thus ?thesis using assms ofilter_Restr_Int[of r A B] by auto -qed - - -lemma ofilter_subset_embed: -assumes WELL: "Well_order r" and - OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" -shows "(A \ B) = (embed (Restr r A) (Restr r B) id)" -proof- - let ?rA = "Restr r A" let ?rB = "Restr r B" - have Well: "wo_rel r" unfolding wo_rel_def using WELL . - hence Refl: "Refl r" by (simp add: wo_rel.REFL) - hence FieldA: "Field ?rA = Field r Int A" - using Refl_Field_Restr by blast - have FieldB: "Field ?rB = Field r Int B" - using Refl Refl_Field_Restr by blast - have WellA: "wo_rel ?rA \ Well_order ?rA" using WELL - by (simp add: Well_order_Restr wo_rel_def) - have WellB: "wo_rel ?rB \ Well_order ?rB" using WELL - by (simp add: Well_order_Restr wo_rel_def) - (* Main proof *) - show ?thesis - proof - assume *: "A \ B" - hence "wo_rel.ofilter (Restr r B) A" using assms - by (simp add: ofilter_Restr_subset) - hence "embed (Restr ?rB A) (Restr r B) id" - using WellB ofilter_embed[of "?rB" A] by auto - thus "embed (Restr r A) (Restr r B) id" - using * by (simp add: Restr_subset) - next - assume *: "embed (Restr r A) (Restr r B) id" - {fix a assume **: "a \ A" - hence "a \ Field r" using Well OFA by (auto simp add: wo_rel.ofilter_def) - with ** FieldA have "a \ Field ?rA" by auto - hence "a \ Field ?rB" using * WellA embed_Field[of ?rA ?rB id] by auto - hence "a \ B" using FieldB by auto - } - thus "A \ B" by blast - qed -qed - - -lemma ofilter_subset_embedS_iso: -assumes WELL: "Well_order r" and - OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" -shows "((A < B) = (embedS (Restr r A) (Restr r B) id)) \ - ((A = B) = (iso (Restr r A) (Restr r B) id))" -proof- - let ?rA = "Restr r A" let ?rB = "Restr r B" - have Well: "wo_rel r" unfolding wo_rel_def using WELL . - hence Refl: "Refl r" by (simp add: wo_rel.REFL) - hence "Field ?rA = Field r Int A" - using Refl_Field_Restr by blast - hence FieldA: "Field ?rA = A" using OFA Well - by (auto simp add: wo_rel.ofilter_def) - have "Field ?rB = Field r Int B" - using Refl Refl_Field_Restr by blast - hence FieldB: "Field ?rB = B" using OFB Well - by (auto simp add: wo_rel.ofilter_def) - (* Main proof *) - show ?thesis unfolding embedS_def iso_def - using assms ofilter_subset_embed[of r A B] - FieldA FieldB bij_betw_id_iff[of A B] by auto -qed - - -lemma ofilter_subset_embedS: -assumes WELL: "Well_order r" and - OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" -shows "(A < B) = embedS (Restr r A) (Restr r B) id" -using assms by (simp add: ofilter_subset_embedS_iso) - - -lemma embed_implies_iso_Restr: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r' r f" -shows "iso r' (Restr r (f ` (Field r'))) f" -proof- - let ?A' = "Field r'" - let ?r'' = "Restr r (f ` ?A')" - have 0: "Well_order ?r''" using WELL Well_order_Restr by blast - have 1: "wo_rel.ofilter r (f ` ?A')" using assms embed_Field_ofilter by blast - hence "Field ?r'' = f ` (Field r')" using WELL Field_Restr_ofilter by blast - hence "bij_betw f ?A' (Field ?r'')" - using EMB embed_inj_on WELL' unfolding bij_betw_def by blast - moreover - {have "\a b. (a,b) \ r' \ a \ Field r' \ b \ Field r'" - unfolding Field_def by auto - hence "compat r' ?r'' f" - using assms embed_iff_compat_inj_on_ofilter - unfolding compat_def by blast - } - ultimately show ?thesis using WELL' 0 iso_iff3 by blast -qed - - -subsection {* The strict inclusion on proper ofilters is well-founded *} - - -definition ofilterIncl :: "'a rel \ 'a set rel" -where -"ofilterIncl r \ {(A,B). wo_rel.ofilter r A \ A \ Field r \ - wo_rel.ofilter r B \ B \ Field r \ A < B}" - - -lemma wf_ofilterIncl: -assumes WELL: "Well_order r" -shows "wf(ofilterIncl r)" -proof- - have Well: "wo_rel r" using WELL by (simp add: wo_rel_def) - hence Lo: "Linear_order r" by (simp add: wo_rel.LIN) - let ?h = "(\ A. wo_rel.suc r A)" - let ?rS = "r - Id" - have "wf ?rS" using WELL by (simp add: order_on_defs) - moreover - have "compat (ofilterIncl r) ?rS ?h" - proof(unfold compat_def ofilterIncl_def, - intro allI impI, simp, elim conjE) - fix A B - assume *: "wo_rel.ofilter r A" "A \ Field r" and - **: "wo_rel.ofilter r B" "B \ Field r" and ***: "A < B" - then obtain a and b where 0: "a \ Field r \ b \ Field r" and - 1: "A = rel.underS r a \ B = rel.underS r b" - using Well by (auto simp add: wo_rel.ofilter_underS_Field) - hence "a \ b" using *** by auto - moreover - have "(a,b) \ r" using 0 1 Lo *** - by (auto simp add: rel.underS_incl_iff) - moreover - have "a = wo_rel.suc r A \ b = wo_rel.suc r B" - using Well 0 1 by (simp add: wo_rel.suc_underS) - ultimately - show "(wo_rel.suc r A, wo_rel.suc r B) \ r \ wo_rel.suc r A \ wo_rel.suc r B" - by simp - qed - ultimately show "wf (ofilterIncl r)" by (simp add: compat_wf) -qed - - - -subsection {* Ordering the well-orders by existence of embeddings *} - - -text {* We define three relations between well-orders: -\begin{itemize} -\item @{text "ordLeq"}, of being embedded (abbreviated @{text "\o"}); -\item @{text "ordLess"}, of being strictly embedded (abbreviated @{text ""}, @{text "<"}, @{text "="} associated to a total order on a set. -*} - - -definition ordLeq :: "('a rel * 'a' rel) set" -where -"ordLeq = {(r,r'). Well_order r \ Well_order r' \ (\f. embed r r' f)}" - - -abbreviation ordLeq2 :: "'a rel \ 'a' rel \ bool" (infix "<=o" 50) -where "r <=o r' \ (r,r') \ ordLeq" - - -abbreviation ordLeq3 :: "'a rel \ 'a' rel \ bool" (infix "\o" 50) -where "r \o r' \ r <=o r'" - - -definition ordLess :: "('a rel * 'a' rel) set" -where -"ordLess = {(r,r'). Well_order r \ Well_order r' \ (\f. embedS r r' f)}" - -abbreviation ordLess2 :: "'a rel \ 'a' rel \ bool" (infix " (r,r') \ ordLess" - - -definition ordIso :: "('a rel * 'a' rel) set" -where -"ordIso = {(r,r'). Well_order r \ Well_order r' \ (\f. iso r r' f)}" - -abbreviation ordIso2 :: "'a rel \ 'a' rel \ bool" (infix "=o" 50) -where "r =o r' \ (r,r') \ ordIso" - - -lemmas ordRels_def = ordLeq_def ordLess_def ordIso_def - -lemma ordLeq_Well_order_simp: -assumes "r \o r'" -shows "Well_order r \ Well_order r'" -using assms unfolding ordLeq_def by simp - - -text{* Notice that the relations @{text "\o"}, @{text " r \o r" -unfolding ordLeq_def using id_embed[of r] by blast - - -lemma ordLeq_transitive[trans]: -assumes *: "r \o r'" and **: "r' \o r''" -shows "r \o r''" -proof- - obtain f and f' - where 1: "Well_order r \ Well_order r' \ Well_order r''" and - "embed r r' f" and "embed r' r'' f'" - using * ** unfolding ordLeq_def by blast - hence "embed r r'' (f' o f)" - using comp_embed[of r r' f r'' f'] by auto - thus "r \o r''" unfolding ordLeq_def using 1 by auto -qed - - -lemma ordLeq_total: -"\Well_order r; Well_order r'\ \ r \o r' \ r' \o r" -unfolding ordLeq_def using wellorders_totally_ordered by blast - - -lemma ordIso_reflexive: -"Well_order r \ r =o r" -unfolding ordIso_def using id_iso[of r] by blast - - -lemma ordIso_transitive[trans]: -assumes *: "r =o r'" and **: "r' =o r''" -shows "r =o r''" -proof- - obtain f and f' - where 1: "Well_order r \ Well_order r' \ Well_order r''" and - "iso r r' f" and 3: "iso r' r'' f'" - using * ** unfolding ordIso_def by auto - hence "iso r r'' (f' o f)" - using comp_iso[of r r' f r'' f'] by auto - thus "r =o r''" unfolding ordIso_def using 1 by auto -qed - - -lemma ordIso_symmetric: -assumes *: "r =o r'" -shows "r' =o r" -proof- - obtain f where 1: "Well_order r \ Well_order r'" and - 2: "embed r r' f \ bij_betw f (Field r) (Field r')" - using * by (auto simp add: ordIso_def iso_def) - let ?f' = "inv_into (Field r) f" - have "embed r' r ?f' \ bij_betw ?f' (Field r') (Field r)" - using 1 2 by (simp add: bij_betw_inv_into inv_into_Field_embed_bij_betw) - thus "r' =o r" unfolding ordIso_def using 1 by (auto simp add: iso_def) -qed - - -lemma ordLeq_ordLess_trans[trans]: -assumes "r \o r'" and " r' Well_order r''" - using assms unfolding ordLeq_def ordLess_def by auto - thus ?thesis using assms unfolding ordLeq_def ordLess_def - using embed_comp_embedS by blast -qed - - -lemma ordLess_ordLeq_trans[trans]: -assumes "r o r''" -shows "r Well_order r''" - using assms unfolding ordLeq_def ordLess_def by auto - thus ?thesis using assms unfolding ordLeq_def ordLess_def - using embedS_comp_embed by blast -qed - - -lemma ordLeq_ordIso_trans[trans]: -assumes "r \o r'" and " r' =o r''" -shows "r \o r''" -proof- - have "Well_order r \ Well_order r''" - using assms unfolding ordLeq_def ordIso_def by auto - thus ?thesis using assms unfolding ordLeq_def ordIso_def - using embed_comp_iso by blast -qed - - -lemma ordIso_ordLeq_trans[trans]: -assumes "r =o r'" and " r' \o r''" -shows "r \o r''" -proof- - have "Well_order r \ Well_order r''" - using assms unfolding ordLeq_def ordIso_def by auto - thus ?thesis using assms unfolding ordLeq_def ordIso_def - using iso_comp_embed by blast -qed - - -lemma ordLess_ordIso_trans[trans]: -assumes "r Well_order r''" - using assms unfolding ordLess_def ordIso_def by auto - thus ?thesis using assms unfolding ordLess_def ordIso_def - using embedS_comp_iso by blast -qed - - -lemma ordIso_ordLess_trans[trans]: -assumes "r =o r'" and " r' Well_order r''" - using assms unfolding ordLess_def ordIso_def by auto - thus ?thesis using assms unfolding ordLess_def ordIso_def - using iso_comp_embedS by blast -qed - - -lemma ordLess_not_embed: -assumes "r (\f'. embed r' r f')" -proof- - obtain f where 1: "Well_order r \ Well_order r'" and 2: "embed r r' f" and - 3: " \ bij_betw f (Field r) (Field r')" - using assms unfolding ordLess_def by (auto simp add: embedS_def) - {fix f' assume *: "embed r' r f'" - hence "bij_betw f (Field r) (Field r')" using 1 2 - by (simp add: embed_bothWays_Field_bij_betw) - with 3 have False by contradiction - } - thus ?thesis by blast -qed - - -lemma ordLess_Field: -assumes OL: "r1 (f`(Field r1) = Field r2)" -proof- - let ?A1 = "Field r1" let ?A2 = "Field r2" - obtain g where - 0: "Well_order r1 \ Well_order r2" and - 1: "embed r1 r2 g \ \(bij_betw g ?A1 ?A2)" - using OL unfolding ordLess_def by (auto simp add: embedS_def) - hence "\a \ ?A1. f a = g a" - using 0 EMB embed_unique[of r1] by auto - hence "\(bij_betw f ?A1 ?A2)" - using 1 bij_betw_cong[of ?A1] by blast - moreover - have "inj_on f ?A1" using EMB 0 by (simp add: embed_inj_on) - ultimately show ?thesis by (simp add: bij_betw_def) -qed - - -lemma ordLess_iff: -"r Well_order r' \ \(\f'. embed r' r f'))" -proof - assume *: "r (\f'. embed r' r f')" using ordLess_not_embed[of r r'] by simp - with * show "Well_order r \ Well_order r' \ \ (\f'. embed r' r f')" - unfolding ordLess_def by auto -next - assume *: "Well_order r \ Well_order r' \ \ (\f'. embed r' r f')" - then obtain f where 1: "embed r r' f" - using wellorders_totally_ordered[of r r'] by blast - moreover - {assume "bij_betw f (Field r) (Field r')" - with * 1 have "embed r' r (inv_into (Field r) f) " - using inv_into_Field_embed_bij_betw[of r r' f] by auto - with * have False by blast - } - ultimately show "(r,r') \ ordLess" - unfolding ordLess_def using * by (fastforce simp add: embedS_def) -qed - - -lemma ordLess_irreflexive: "\ r \(\f. embed r r f)" - unfolding ordLess_iff .. - moreover have "embed r r id" using id_embed[of r] . - ultimately show False by blast -qed - - -lemma ordLeq_iff_ordLess_or_ordIso: -"r \o r' = (r r =o r')" -unfolding ordRels_def embedS_defs iso_defs by blast - - -lemma ordIso_iff_ordLeq: -"(r =o r') = (r \o r' \ r' \o r)" -proof - assume "r =o r'" - then obtain f where 1: "Well_order r \ Well_order r' \ - embed r r' f \ bij_betw f (Field r) (Field r')" - unfolding ordIso_def iso_defs by auto - hence "embed r r' f \ embed r' r (inv_into (Field r) f)" - by (simp add: inv_into_Field_embed_bij_betw) - thus "r \o r' \ r' \o r" - unfolding ordLeq_def using 1 by auto -next - assume "r \o r' \ r' \o r" - then obtain f and g where 1: "Well_order r \ Well_order r' \ - embed r r' f \ embed r' r g" - unfolding ordLeq_def by auto - hence "iso r r' f" by (auto simp add: embed_bothWays_iso) - thus "r =o r'" unfolding ordIso_def using 1 by auto -qed - - -lemma not_ordLess_ordLeq: -"r \ r' \o r" -using ordLess_ordLeq_trans ordLess_irreflexive by blast - - -lemma ordLess_or_ordLeq: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "r r' \o r" -proof- - have "r \o r' \ r' \o r" - using assms by (simp add: ordLeq_total) - moreover - {assume "\ r r \o r'" - hence "r =o r'" using ordLeq_iff_ordLess_or_ordIso by blast - hence "r' \o r" using ordIso_symmetric ordIso_iff_ordLeq by blast - } - ultimately show ?thesis by blast -qed - - -lemma not_ordLess_ordIso: -"r \ r =o r'" -using assms ordLess_ordIso_trans ordIso_symmetric ordLess_irreflexive by blast - - -lemma not_ordLeq_iff_ordLess: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "(\ r' \o r) = (r r' o r')" -using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast - - -lemma ordLess_transitive[trans]: -"\r \ r r \o r'" -using ordIso_iff_ordLeq by blast - - -lemma ordLess_imp_ordLeq: -"r r \o r'" -using ordLeq_iff_ordLess_or_ordIso by blast - - -lemma ofilter_subset_ordLeq: -assumes WELL: "Well_order r" and - OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" -shows "(A \ B) = (Restr r A \o Restr r B)" -proof - assume "A \ B" - thus "Restr r A \o Restr r B" - unfolding ordLeq_def using assms - Well_order_Restr Well_order_Restr ofilter_subset_embed by blast -next - assume *: "Restr r A \o Restr r B" - then obtain f where "embed (Restr r A) (Restr r B) f" - unfolding ordLeq_def by blast - {assume "B < A" - hence "Restr r B B" using OFA OFB WELL - wo_rel_def[of r] wo_rel.ofilter_linord[of r A B] by blast -qed - - -lemma ofilter_subset_ordLess: -assumes WELL: "Well_order r" and - OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" -shows "(A < B) = (Restr r A Well_order ?rB" - using WELL Well_order_Restr by blast - have "(A < B) = (\ B \ A)" using assms - wo_rel_def wo_rel.ofilter_linord[of r A B] by blast - also have "\ = (\ Restr r B \o Restr r A)" - using assms ofilter_subset_ordLeq by blast - also have "\ = (Restr r A Well_order r; wo_rel.ofilter r A\ \ (A < Field r) = (Restr r A {}" -shows "Restr r (rel.underS r a) (ofilterIncl r3)" -proof- - have OL13: "r1 Well_order r2 \ Well_order r3" and - 1: "embed r1 r2 f12 \ \(bij_betw f12 ?A1 ?A2)" and - 2: "embed r2 r3 g23 \ \(bij_betw g23 ?A2 ?A3)" - using OL12 OL23 by (auto simp add: ordLess_def embedS_def) - hence "\a \ ?A2. f23 a = g23 a" - using EMB23 embed_unique[of r2 r3] by blast - hence 3: "\(bij_betw f23 ?A2 ?A3)" - using 2 bij_betw_cong[of ?A2 f23 g23] by blast - (* *) - have 4: "wo_rel.ofilter r2 (f12 ` ?A1) \ f12 ` ?A1 \ ?A2" - using 0 1 OL12 by (simp add: embed_Field_ofilter ordLess_Field) - have 5: "wo_rel.ofilter r3 (f23 ` ?A2) \ f23 ` ?A2 \ ?A3" - using 0 EMB23 OL23 by (simp add: embed_Field_ofilter ordLess_Field) - have 6: "wo_rel.ofilter r3 (f13 ` ?A1) \ f13 ` ?A1 \ ?A3" - using 0 EMB13 OL13 by (simp add: embed_Field_ofilter ordLess_Field) - (* *) - have "f12 ` ?A1 < ?A2" - using 0 4 by (auto simp add: wo_rel_def wo_rel.ofilter_def) - moreover have "inj_on f23 ?A2" - using EMB23 0 by (simp add: wo_rel_def embed_inj_on) - ultimately - have "f23 ` (f12 ` ?A1) < f23 ` ?A2" by (simp add: inj_on_strict_subset) - moreover - {have "embed r1 r3 (f23 o f12)" - using 1 EMB23 0 by (auto simp add: comp_embed) - hence "\a \ ?A1. f23(f12 a) = f13 a" - using EMB13 0 embed_unique[of r1 r3 "f23 o f12" f13] by auto - hence "f23 ` (f12 ` ?A1) = f13 ` ?A1" by force - } - ultimately - have "f13 ` ?A1 < f23 ` ?A2" by simp - (* *) - with 5 6 show ?thesis - unfolding ofilterIncl_def by auto -qed - - -lemma ordLess_iff_ordIso_Restr: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "(r' a \ Field r. r' =o Restr r (rel.underS r a))" -proof(auto) - fix a assume *: "a \ Field r" and **: "r' =o Restr r (rel.underS r a)" - hence "Restr r (rel.underS r a) Well_order r'" and - 2: "embed r' r f \ f ` (Field r') \ Field r" - unfolding ordLess_def embedS_def[abs_def] bij_betw_def using embed_inj_on by blast - hence "wo_rel.ofilter r (f ` (Field r'))" using embed_Field_ofilter by blast - then obtain a where 3: "a \ Field r" and 4: "rel.underS r a = f ` (Field r')" - using 1 2 by (auto simp add: wo_rel.ofilter_underS_Field wo_rel_def) - have "iso r' (Restr r (f ` (Field r'))) f" - using embed_implies_iso_Restr 2 assms by blast - moreover have "Well_order (Restr r (f ` (Field r')))" - using WELL Well_order_Restr by blast - ultimately have "r' =o Restr r (f ` (Field r'))" - using WELL' unfolding ordIso_def by auto - hence "r' =o Restr r (rel.underS r a)" using 4 by auto - thus "\a \ Field r. r' =o Restr r (rel.underS r a)" using 3 by auto -qed - - -lemma internalize_ordLess: -"(r' p. Field p < Field r \ r' =o p \ p Well_order r'" unfolding ordLess_def by auto - with * obtain a where 1: "a \ Field r" and 2: "r' =o Restr r (rel.underS r a)" - using ordLess_iff_ordIso_Restr by blast - let ?p = "Restr r (rel.underS r a)" - have "wo_rel.ofilter r (rel.underS r a)" using 0 - by (simp add: wo_rel_def wo_rel.underS_ofilter) - hence "Field ?p = rel.underS r a" using 0 Field_Restr_ofilter by blast - hence "Field ?p < Field r" using rel.underS_Field2 1 by fastforce - moreover have "?p p. Field p < Field r \ r' =o p \ p p. Field p < Field r \ r' =o p \ p o r) = (\p. Field p \ Field r \ r' =o p \ p \o r)" -proof - assume *: "r' \o r" - moreover - {assume "r' r' =o p \ p p. Field p \ Field r \ r' =o p \ p \o r" - using ordLeq_iff_ordLess_or_ordIso by blast - } - moreover - have "r \o r" using * ordLeq_def ordLeq_reflexive by blast - ultimately show "\p. Field p \ Field r \ r' =o p \ p \o r" - using ordLeq_iff_ordLess_or_ordIso by blast -next - assume "\p. Field p \ Field r \ r' =o p \ p \o r" - thus "r' \o r" using ordIso_ordLeq_trans by blast -qed - - -lemma ordLeq_iff_ordLess_Restr: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "(r \o r') = (\a \ Field r. Restr r (rel.underS r a) o r'" - fix a assume "a \ Field r" - hence "Restr r (rel.underS r a) a \ Field r. Restr r (rel.underS r a) Field r \ r' =o Restr r (rel.underS r a)" - using assms ordLess_iff_ordIso_Restr by blast - hence False using * not_ordLess_ordIso ordIso_symmetric by blast - } - thus "r \o r'" using ordLess_or_ordLeq assms by blast -qed - - -lemma finite_ordLess_infinite: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - FIN: "finite(Field r)" and INF: "infinite(Field r')" -shows "r o r" - then obtain h where "inj_on h (Field r') \ h ` (Field r') \ Field r" - unfolding ordLeq_def using assms embed_inj_on embed_Field by blast - hence False using finite_imageD finite_subset FIN INF by metis - } - thus ?thesis using WELL WELL' ordLess_or_ordLeq by blast -qed - - -lemma finite_well_order_on_ordIso: -assumes FIN: "finite A" and - WELL: "well_order_on A r" and WELL': "well_order_on A r'" -shows "r =o r'" -proof- - have 0: "Well_order r \ Well_order r' \ Field r = A \ Field r' = A" - using assms rel.well_order_on_Well_order by blast - moreover - have "\r r'. well_order_on A r \ well_order_on A r' \ r \o r' - \ r =o r'" - proof(clarify) - fix r r' assume *: "well_order_on A r" and **: "well_order_on A r'" - have 2: "Well_order r \ Well_order r' \ Field r = A \ Field r' = A" - using * ** rel.well_order_on_Well_order by blast - assume "r \o r'" - then obtain f where 1: "embed r r' f" and - "inj_on f A \ f ` A \ A" - unfolding ordLeq_def using 2 embed_inj_on embed_Field by blast - hence "bij_betw f A A" unfolding bij_betw_def using FIN endo_inj_surj by blast - thus "r =o r'" unfolding ordIso_def iso_def[abs_def] using 1 2 by auto - qed - ultimately show ?thesis using assms ordLeq_total ordIso_symmetric by blast -qed - - -subsection{* @{text " 'a rel \ 'a set" -where "ord_to_filter r0 r \ (SOME f. embed r r0 f) ` (Field r)" - - -lemma ord_to_filter_compat: -"compat (ordLess Int (ordLess^-1``{r0} \ ordLess^-1``{r0})) - (ofilterIncl r0) - (ord_to_filter r0)" -proof(unfold compat_def ord_to_filter_def, clarify) - fix r1::"'a rel" and r2::"'a rel" - let ?A1 = "Field r1" let ?A2 ="Field r2" let ?A0 ="Field r0" - let ?phi10 = "\ f10. embed r1 r0 f10" let ?f10 = "SOME f. ?phi10 f" - let ?phi20 = "\ f20. embed r2 r0 f20" let ?f20 = "SOME f. ?phi20 f" - assume *: "r1 f. ?phi10 f) \ (\f. ?phi20 f)" - by (auto simp add: ordLess_def embedS_def) - hence "?phi10 ?f10 \ ?phi20 ?f20" by (auto simp add: someI_ex) - thus "(?f10 ` ?A1, ?f20 ` ?A2) \ ofilterIncl r0" - using * ** by (simp add: embed_ordLess_ofilterIncl) -qed - - -theorem wf_ordLess: "wf ordLess" -proof- - {fix r0 :: "('a \ 'a) set" - (* need to annotate here!*) - let ?ordLess = "ordLess::('d rel * 'd rel) set" - let ?R = "?ordLess Int (?ordLess^-1``{r0} \ ?ordLess^-1``{r0})" - {assume Case1: "Well_order r0" - hence "wf ?R" - using wf_ofilterIncl[of r0] - compat_wf[of ?R "ofilterIncl r0" "ord_to_filter r0"] - ord_to_filter_compat[of r0] by auto - } - moreover - {assume Case2: "\ Well_order r0" - hence "?R = {}" unfolding ordLess_def by auto - hence "wf ?R" using wf_empty by simp - } - ultimately have "wf ?R" by blast - } - thus ?thesis by (simp add: trans_wf_iff ordLess_trans) -qed - -corollary exists_minim_Well_order: -assumes NE: "R \ {}" and WELL: "\r \ R. Well_order r" -shows "\r \ R. \r' \ R. r \o r'" -proof- - obtain r where "r \ R \ (\r' \ R. \ r' ('a \ 'a') \ 'a' rel" -where -"dir_image r f = {(f a, f b)| a b. (a,b) \ r}" - - -lemma dir_image_Field: -"Field(dir_image r f) \ f ` (Field r)" -unfolding dir_image_def Field_def by auto - - -lemma dir_image_minus_Id: -"inj_on f (Field r) \ (dir_image r f) - Id = dir_image (r - Id) f" -unfolding inj_on_def Field_def dir_image_def by auto - - -lemma Refl_dir_image: -assumes "Refl r" -shows "Refl(dir_image r f)" -proof- - {fix a' b' - assume "(a',b') \ dir_image r f" - then obtain a b where 1: "a' = f a \ b' = f b \ (a,b) \ r" - unfolding dir_image_def by blast - hence "a \ Field r \ b \ Field r" using Field_def by fastforce - hence "(a,a) \ r \ (b,b) \ r" using assms by (simp add: refl_on_def) - with 1 have "(a',a') \ dir_image r f \ (b',b') \ dir_image r f" - unfolding dir_image_def by auto - } - thus ?thesis - by(unfold refl_on_def Field_def Domain_def Range_def, auto) -qed - - -lemma trans_dir_image: -assumes TRANS: "trans r" and INJ: "inj_on f (Field r)" -shows "trans(dir_image r f)" -proof(unfold trans_def, auto) - fix a' b' c' - assume "(a',b') \ dir_image r f" "(b',c') \ dir_image r f" - then obtain a b1 b2 c where 1: "a' = f a \ b' = f b1 \ b' = f b2 \ c' = f c" and - 2: "(a,b1) \ r \ (b2,c) \ r" - unfolding dir_image_def by blast - hence "b1 \ Field r \ b2 \ Field r" - unfolding Field_def by auto - hence "b1 = b2" using 1 INJ unfolding inj_on_def by auto - hence "(a,c): r" using 2 TRANS unfolding trans_def by blast - thus "(a',c') \ dir_image r f" - unfolding dir_image_def using 1 by auto -qed - - -lemma Preorder_dir_image: -"\Preorder r; inj_on f (Field r)\ \ Preorder (dir_image r f)" -by (simp add: preorder_on_def Refl_dir_image trans_dir_image) - - -lemma antisym_dir_image: -assumes AN: "antisym r" and INJ: "inj_on f (Field r)" -shows "antisym(dir_image r f)" -proof(unfold antisym_def, auto) - fix a' b' - assume "(a',b') \ dir_image r f" "(b',a') \ dir_image r f" - then obtain a1 b1 a2 b2 where 1: "a' = f a1 \ a' = f a2 \ b' = f b1 \ b' = f b2" and - 2: "(a1,b1) \ r \ (b2,a2) \ r " and - 3: "{a1,a2,b1,b2} \ Field r" - unfolding dir_image_def Field_def by blast - hence "a1 = a2 \ b1 = b2" using INJ unfolding inj_on_def by auto - hence "a1 = b2" using 2 AN unfolding antisym_def by auto - thus "a' = b'" using 1 by auto -qed - - -lemma Partial_order_dir_image: -"\Partial_order r; inj_on f (Field r)\ \ Partial_order (dir_image r f)" -by (simp add: partial_order_on_def Preorder_dir_image antisym_dir_image) - - -lemma Total_dir_image: -assumes TOT: "Total r" and INJ: "inj_on f (Field r)" -shows "Total(dir_image r f)" -proof(unfold total_on_def, intro ballI impI) - fix a' b' - assume "a' \ Field (dir_image r f)" "b' \ Field (dir_image r f)" - then obtain a and b where 1: "a \ Field r \ b \ Field r \ f a = a' \ f b = b'" - using dir_image_Field[of r f] by blast - moreover assume "a' \ b'" - ultimately have "a \ b" using INJ unfolding inj_on_def by auto - hence "(a,b) \ r \ (b,a) \ r" using 1 TOT unfolding total_on_def by auto - thus "(a',b') \ dir_image r f \ (b',a') \ dir_image r f" - using 1 unfolding dir_image_def by auto -qed - - -lemma Linear_order_dir_image: -"\Linear_order r; inj_on f (Field r)\ \ Linear_order (dir_image r f)" -by (simp add: linear_order_on_def Partial_order_dir_image Total_dir_image) - - -lemma wf_dir_image: -assumes WF: "wf r" and INJ: "inj_on f (Field r)" -shows "wf(dir_image r f)" -proof(unfold wf_eq_minimal2, intro allI impI, elim conjE) - fix A'::"'b set" - assume SUB: "A' \ Field(dir_image r f)" and NE: "A' \ {}" - obtain A where A_def: "A = {a \ Field r. f a \ A'}" by blast - have "A \ {} \ A \ Field r" - using A_def dir_image_Field[of r f] SUB NE by blast - then obtain a where 1: "a \ A \ (\b \ A. (b,a) \ r)" - using WF unfolding wf_eq_minimal2 by metis - have "\b' \ A'. (b',f a) \ dir_image r f" - proof(clarify) - fix b' assume *: "b' \ A'" and **: "(b',f a) \ dir_image r f" - obtain b1 a1 where 2: "b' = f b1 \ f a = f a1" and - 3: "(b1,a1) \ r \ {a1,b1} \ Field r" - using ** unfolding dir_image_def Field_def by blast - hence "a = a1" using 1 A_def INJ unfolding inj_on_def by auto - hence "b1 \ A \ (b1,a) \ r" using 2 3 A_def * by auto - with 1 show False by auto - qed - thus "\a'\A'. \b'\A'. (b', a') \ dir_image r f" - using A_def 1 by blast -qed - - -lemma Well_order_dir_image: -"\Well_order r; inj_on f (Field r)\ \ Well_order (dir_image r f)" -using assms unfolding well_order_on_def -using Linear_order_dir_image[of r f] wf_dir_image[of "r - Id" f] - dir_image_minus_Id[of f r] - subset_inj_on[of f "Field r" "Field(r - Id)"] - mono_Field[of "r - Id" r] by auto - - -lemma dir_image_Field2: -"Refl r \ Field(dir_image r f) = f ` (Field r)" -unfolding Field_def dir_image_def refl_on_def Domain_def Range_def by blast - - -lemma dir_image_bij_betw: -"\Well_order r; inj_on f (Field r)\ \ bij_betw f (Field r) (Field (dir_image r f))" -unfolding bij_betw_def -by (simp add: dir_image_Field2 order_on_defs) - - -lemma dir_image_compat: -"compat r (dir_image r f) f" -unfolding compat_def dir_image_def by auto - - -lemma dir_image_iso: -"\Well_order r; inj_on f (Field r)\ \ iso r (dir_image r f) f" -using iso_iff3 dir_image_compat dir_image_bij_betw Well_order_dir_image by blast - - -lemma dir_image_ordIso: -"\Well_order r; inj_on f (Field r)\ \ r =o dir_image r f" -unfolding ordIso_def using dir_image_iso Well_order_dir_image by blast - - -lemma Well_order_iso_copy: -assumes WELL: "well_order_on A r" and BIJ: "bij_betw f A A'" -shows "\r'. well_order_on A' r' \ r =o r'" -proof- - let ?r' = "dir_image r f" - have 1: "A = Field r \ Well_order r" - using WELL rel.well_order_on_Well_order by blast - hence 2: "iso r ?r' f" - using dir_image_iso using BIJ unfolding bij_betw_def by auto - hence "f ` (Field r) = Field ?r'" using 1 iso_iff[of r ?r'] by blast - hence "Field ?r' = A'" - using 1 BIJ unfolding bij_betw_def by auto - moreover have "Well_order ?r'" - using 1 Well_order_dir_image BIJ unfolding bij_betw_def by blast - ultimately show ?thesis unfolding ordIso_def using 1 2 by blast -qed - - - -subsection {* Bounded square *} - - -text{* This construction essentially defines, for an order relation @{text "r"}, a lexicographic -order @{text "bsqr r"} on @{text "(Field r) \ (Field r)"}, applying the -following criteria (in this order): -\begin{itemize} -\item compare the maximums; -\item compare the first components; -\item compare the second components. -\end{itemize} -% -The only application of this construction that we are aware of is -at proving that the square of an infinite set has the same cardinal -as that set. The essential property required there (and which is ensured by this -construction) is that any proper order filter of the product order is included in a rectangle, i.e., -in a product of proper filters on the original relation (assumed to be a well-order). *} - - -definition bsqr :: "'a rel => ('a * 'a)rel" -where -"bsqr r = {((a1,a2),(b1,b2)). - {a1,a2,b1,b2} \ Field r \ - (a1 = b1 \ a2 = b2 \ - (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ - wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ - wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id - )}" - - -lemma Field_bsqr: -"Field (bsqr r) = Field r \ Field r" -proof - show "Field (bsqr r) \ Field r \ Field r" - proof- - {fix a1 a2 assume "(a1,a2) \ Field (bsqr r)" - moreover - have "\ b1 b2. ((a1,a2),(b1,b2)) \ bsqr r \ ((b1,b2),(a1,a2)) \ bsqr r \ - a1 \ Field r \ a2 \ Field r" unfolding bsqr_def by auto - ultimately have "a1 \ Field r \ a2 \ Field r" unfolding Field_def by auto - } - thus ?thesis unfolding Field_def by force - qed -next - show "Field r \ Field r \ Field (bsqr r)" - proof(auto) - fix a1 a2 assume "a1 \ Field r" and "a2 \ Field r" - hence "((a1,a2),(a1,a2)) \ bsqr r" unfolding bsqr_def by blast - thus "(a1,a2) \ Field (bsqr r)" unfolding Field_def by auto - qed -qed - - -lemma bsqr_Refl: "Refl(bsqr r)" -by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def) - - -lemma bsqr_Trans: -assumes "Well_order r" -shows "trans (bsqr r)" -proof(unfold trans_def, auto) - (* Preliminary facts *) - have Well: "wo_rel r" using assms wo_rel_def by auto - hence Trans: "trans r" using wo_rel.TRANS by auto - have Anti: "antisym r" using wo_rel.ANTISYM Well by auto - hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id) - (* Main proof *) - fix a1 a2 b1 b2 c1 c2 - assume *: "((a1,a2),(b1,b2)) \ bsqr r" and **: "((b1,b2),(c1,c2)) \ bsqr r" - hence 0: "{a1,a2,b1,b2,c1,c2} \ Field r" unfolding bsqr_def by auto - have 1: "a1 = b1 \ a2 = b2 \ (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ - wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ - wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" - using * unfolding bsqr_def by auto - have 2: "b1 = c1 \ b2 = c2 \ (wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id \ - wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id \ - wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1 \ (b2,c2) \ r - Id" - using ** unfolding bsqr_def by auto - show "((a1,a2),(c1,c2)) \ bsqr r" - proof- - {assume Case1: "a1 = b1 \ a2 = b2" - hence ?thesis using ** by simp - } - moreover - {assume Case2: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id" - {assume Case21: "b1 = c1 \ b2 = c2" - hence ?thesis using * by simp - } - moreover - {assume Case22: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" - hence "(wo_rel.max2 r a1 a2, wo_rel.max2 r c1 c2) \ r - Id" - using Case2 TransS trans_def[of "r - Id"] by blast - hence ?thesis using 0 unfolding bsqr_def by auto - } - moreover - {assume Case23_4: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2" - hence ?thesis using Case2 0 unfolding bsqr_def by auto - } - ultimately have ?thesis using 0 2 by auto - } - moreover - {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id" - {assume Case31: "b1 = c1 \ b2 = c2" - hence ?thesis using * by simp - } - moreover - {assume Case32: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" - hence ?thesis using Case3 0 unfolding bsqr_def by auto - } - moreover - {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id" - hence "(a1,c1) \ r - Id" - using Case3 TransS trans_def[of "r - Id"] by blast - hence ?thesis using Case3 Case33 0 unfolding bsqr_def by auto - } - moreover - {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1" - hence ?thesis using Case3 0 unfolding bsqr_def by auto - } - ultimately have ?thesis using 0 2 by auto - } - moreover - {assume Case4: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" - {assume Case41: "b1 = c1 \ b2 = c2" - hence ?thesis using * by simp - } - moreover - {assume Case42: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" - hence ?thesis using Case4 0 unfolding bsqr_def by force - } - moreover - {assume Case43: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id" - hence ?thesis using Case4 0 unfolding bsqr_def by auto - } - moreover - {assume Case44: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1 \ (b2,c2) \ r - Id" - hence "(a2,c2) \ r - Id" - using Case4 TransS trans_def[of "r - Id"] by blast - hence ?thesis using Case4 Case44 0 unfolding bsqr_def by auto - } - ultimately have ?thesis using 0 2 by auto - } - ultimately show ?thesis using 0 1 by auto - qed -qed - - -lemma bsqr_antisym: -assumes "Well_order r" -shows "antisym (bsqr r)" -proof(unfold antisym_def, clarify) - (* Preliminary facts *) - have Well: "wo_rel r" using assms wo_rel_def by auto - hence Trans: "trans r" using wo_rel.TRANS by auto - have Anti: "antisym r" using wo_rel.ANTISYM Well by auto - hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id) - hence IrrS: "\a b. \((a,b) \ r - Id \ (b,a) \ r - Id)" - using Anti trans_def[of "r - Id"] antisym_def[of "r - Id"] by blast - (* Main proof *) - fix a1 a2 b1 b2 - assume *: "((a1,a2),(b1,b2)) \ bsqr r" and **: "((b1,b2),(a1,a2)) \ bsqr r" - hence 0: "{a1,a2,b1,b2} \ Field r" unfolding bsqr_def by auto - have 1: "a1 = b1 \ a2 = b2 \ (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ - wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ - wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" - using * unfolding bsqr_def by auto - have 2: "b1 = a1 \ b2 = a2 \ (wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id \ - wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \ (b1,a1) \ r - Id \ - wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \ b1 = a1 \ (b2,a2) \ r - Id" - using ** unfolding bsqr_def by auto - show "a1 = b1 \ a2 = b2" - proof- - {assume Case1: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id" - {assume Case11: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" - hence False using Case1 IrrS by blast - } - moreover - {assume Case12_3: "wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2" - hence False using Case1 by auto - } - ultimately have ?thesis using 0 2 by auto - } - moreover - {assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id" - {assume Case21: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" - hence False using Case2 by auto - } - moreover - {assume Case22: "(b1,a1) \ r - Id" - hence False using Case2 IrrS by blast - } - moreover - {assume Case23: "b1 = a1" - hence False using Case2 by auto - } - ultimately have ?thesis using 0 2 by auto - } - moreover - {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" - moreover - {assume Case31: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" - hence False using Case3 by auto - } - moreover - {assume Case32: "(b1,a1) \ r - Id" - hence False using Case3 by auto - } - moreover - {assume Case33: "(b2,a2) \ r - Id" - hence False using Case3 IrrS by blast - } - ultimately have ?thesis using 0 2 by auto - } - ultimately show ?thesis using 0 1 by blast - qed -qed - - -lemma bsqr_Total: -assumes "Well_order r" -shows "Total(bsqr r)" -proof- - (* Preliminary facts *) - have Well: "wo_rel r" using assms wo_rel_def by auto - hence Total: "\a \ Field r. \b \ Field r. (a,b) \ r \ (b,a) \ r" - using wo_rel.TOTALS by auto - (* Main proof *) - {fix a1 a2 b1 b2 assume "{(a1,a2), (b1,b2)} \ Field(bsqr r)" - hence 0: "a1 \ Field r \ a2 \ Field r \ b1 \ Field r \ b2 \ Field r" - using Field_bsqr by blast - have "((a1,a2) = (b1,b2) \ ((a1,a2),(b1,b2)) \ bsqr r \ ((b1,b2),(a1,a2)) \ bsqr r)" - proof(rule wo_rel.cases_Total[of r a1 a2], clarsimp simp add: Well, simp add: 0) - (* Why didn't clarsimp simp add: Well 0 do the same job? *) - assume Case1: "(a1,a2) \ r" - hence 1: "wo_rel.max2 r a1 a2 = a2" - using Well 0 by (simp add: wo_rel.max2_equals2) - show ?thesis - proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0) - assume Case11: "(b1,b2) \ r" - hence 2: "wo_rel.max2 r b1 b2 = b2" - using Well 0 by (simp add: wo_rel.max2_equals2) - show ?thesis - proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) - assume Case111: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" - thus ?thesis using 0 1 2 unfolding bsqr_def by auto - next - assume Case112: "a2 = b2" - show ?thesis - proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) - assume Case1121: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" - thus ?thesis using 0 1 2 Case112 unfolding bsqr_def by auto - next - assume Case1122: "a1 = b1" - thus ?thesis using Case112 by auto - qed - qed - next - assume Case12: "(b2,b1) \ r" - hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1) - show ?thesis - proof(rule wo_rel.cases_Total3[of r a2 b1], clarsimp simp add: Well, simp add: 0) - assume Case121: "(a2,b1) \ r - Id \ (b1,a2) \ r - Id" - thus ?thesis using 0 1 3 unfolding bsqr_def by auto - next - assume Case122: "a2 = b1" - show ?thesis - proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) - assume Case1221: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" - thus ?thesis using 0 1 3 Case122 unfolding bsqr_def by auto - next - assume Case1222: "a1 = b1" - show ?thesis - proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) - assume Case12221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" - thus ?thesis using 0 1 3 Case122 Case1222 unfolding bsqr_def by auto - next - assume Case12222: "a2 = b2" - thus ?thesis using Case122 Case1222 by auto - qed - qed - qed - qed - next - assume Case2: "(a2,a1) \ r" - hence 1: "wo_rel.max2 r a1 a2 = a1" using Well 0 by (simp add: wo_rel.max2_equals1) - show ?thesis - proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0) - assume Case21: "(b1,b2) \ r" - hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2) - show ?thesis - proof(rule wo_rel.cases_Total3[of r a1 b2], clarsimp simp add: Well, simp add: 0) - assume Case211: "(a1,b2) \ r - Id \ (b2,a1) \ r - Id" - thus ?thesis using 0 1 2 unfolding bsqr_def by auto - next - assume Case212: "a1 = b2" - show ?thesis - proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) - assume Case2121: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" - thus ?thesis using 0 1 2 Case212 unfolding bsqr_def by auto - next - assume Case2122: "a1 = b1" - show ?thesis - proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) - assume Case21221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" - thus ?thesis using 0 1 2 Case212 Case2122 unfolding bsqr_def by auto - next - assume Case21222: "a2 = b2" - thus ?thesis using Case2122 Case212 by auto - qed - qed - qed - next - assume Case22: "(b2,b1) \ r" - hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1) - show ?thesis - proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) - assume Case221: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" - thus ?thesis using 0 1 3 unfolding bsqr_def by auto - next - assume Case222: "a1 = b1" - show ?thesis - proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) - assume Case2221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" - thus ?thesis using 0 1 3 Case222 unfolding bsqr_def by auto - next - assume Case2222: "a2 = b2" - thus ?thesis using Case222 by auto - qed - qed - qed - qed - } - thus ?thesis unfolding total_on_def by fast -qed - - -lemma bsqr_Linear_order: -assumes "Well_order r" -shows "Linear_order(bsqr r)" -unfolding order_on_defs -using assms bsqr_Refl bsqr_Trans bsqr_antisym bsqr_Total by blast - - -lemma bsqr_Well_order: -assumes "Well_order r" -shows "Well_order(bsqr r)" -using assms -proof(simp add: bsqr_Linear_order Linear_order_Well_order_iff, intro allI impI) - have 0: "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r)" - using assms well_order_on_def Linear_order_Well_order_iff by blast - fix D assume *: "D \ Field (bsqr r)" and **: "D \ {}" - hence 1: "D \ Field r \ Field r" unfolding Field_bsqr by simp - (* *) - obtain M where M_def: "M = {wo_rel.max2 r a1 a2| a1 a2. (a1,a2) \ D}" by blast - have "M \ {}" using 1 M_def ** by auto - moreover - have "M \ Field r" unfolding M_def - using 1 assms wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce - ultimately obtain m where m_min: "m \ M \ (\a \ M. (m,a) \ r)" - using 0 by blast - (* *) - obtain A1 where A1_def: "A1 = {a1. \a2. (a1,a2) \ D \ wo_rel.max2 r a1 a2 = m}" by blast - have "A1 \ Field r" unfolding A1_def using 1 by auto - moreover have "A1 \ {}" unfolding A1_def using m_min unfolding M_def by blast - ultimately obtain a1 where a1_min: "a1 \ A1 \ (\a \ A1. (a1,a) \ r)" - using 0 by blast - (* *) - obtain A2 where A2_def: "A2 = {a2. (a1,a2) \ D \ wo_rel.max2 r a1 a2 = m}" by blast - have "A2 \ Field r" unfolding A2_def using 1 by auto - moreover have "A2 \ {}" unfolding A2_def - using m_min a1_min unfolding A1_def M_def by blast - ultimately obtain a2 where a2_min: "a2 \ A2 \ (\a \ A2. (a2,a) \ r)" - using 0 by blast - (* *) - have 2: "wo_rel.max2 r a1 a2 = m" - using a1_min a2_min unfolding A1_def A2_def by auto - have 3: "(a1,a2) \ D" using a2_min unfolding A2_def by auto - (* *) - moreover - {fix b1 b2 assume ***: "(b1,b2) \ D" - hence 4: "{a1,a2,b1,b2} \ Field r" using 1 3 by blast - have 5: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r" - using *** a1_min a2_min m_min unfolding A1_def A2_def M_def by auto - have "((a1,a2),(b1,b2)) \ bsqr r" - proof(cases "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2") - assume Case1: "wo_rel.max2 r a1 a2 \ wo_rel.max2 r b1 b2" - thus ?thesis unfolding bsqr_def using 4 5 by auto - next - assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2" - hence "b1 \ A1" unfolding A1_def using 2 *** by auto - hence 6: "(a1,b1) \ r" using a1_min by auto - show ?thesis - proof(cases "a1 = b1") - assume Case21: "a1 \ b1" - thus ?thesis unfolding bsqr_def using 4 Case2 6 by auto - next - assume Case22: "a1 = b1" - hence "b2 \ A2" unfolding A2_def using 2 *** Case2 by auto - hence 7: "(a2,b2) \ r" using a2_min by auto - thus ?thesis unfolding bsqr_def using 4 7 Case2 Case22 by auto - qed - qed - } - (* *) - ultimately show "\d \ D. \d' \ D. (d,d') \ bsqr r" by fastforce -qed - - -lemma bsqr_max2: -assumes WELL: "Well_order r" and LEQ: "((a1,a2),(b1,b2)) \ bsqr r" -shows "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r" -proof- - have "{(a1,a2),(b1,b2)} \ Field(bsqr r)" - using LEQ unfolding Field_def by auto - hence "{a1,a2,b1,b2} \ Field r" unfolding Field_bsqr by auto - hence "{wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2} \ Field r" - using WELL wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce - moreover have "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2" - using LEQ unfolding bsqr_def by auto - ultimately show ?thesis using WELL unfolding order_on_defs refl_on_def by auto -qed - - -lemma bsqr_ofilter: -assumes WELL: "Well_order r" and - OF: "wo_rel.ofilter (bsqr r) D" and SUB: "D < Field r \ Field r" and - NE: "\ (\a. Field r = rel.under r a)" -shows "\A. wo_rel.ofilter r A \ A < Field r \ D \ A \ A" -proof- - let ?r' = "bsqr r" - have Well: "wo_rel r" using WELL wo_rel_def by blast - hence Trans: "trans r" using wo_rel.TRANS by blast - have Well': "Well_order ?r' \ wo_rel ?r'" - using WELL bsqr_Well_order wo_rel_def by blast - (* *) - have "D < Field ?r'" unfolding Field_bsqr using SUB . - with OF obtain a1 and a2 where - "(a1,a2) \ Field ?r'" and 1: "D = rel.underS ?r' (a1,a2)" - using Well' wo_rel.ofilter_underS_Field[of ?r' D] by auto - hence 2: "{a1,a2} \ Field r" unfolding Field_bsqr by auto - let ?m = "wo_rel.max2 r a1 a2" - have "D \ (rel.under r ?m) \ (rel.under r ?m)" - proof(unfold 1) - {fix b1 b2 - let ?n = "wo_rel.max2 r b1 b2" - assume "(b1,b2) \ rel.underS ?r' (a1,a2)" - hence 3: "((b1,b2),(a1,a2)) \ ?r'" - unfolding rel.underS_def by blast - hence "(?n,?m) \ r" using WELL by (simp add: bsqr_max2) - moreover - {have "(b1,b2) \ Field ?r'" using 3 unfolding Field_def by auto - hence "{b1,b2} \ Field r" unfolding Field_bsqr by auto - hence "(b1,?n) \ r \ (b2,?n) \ r" - using Well by (simp add: wo_rel.max2_greater) - } - ultimately have "(b1,?m) \ r \ (b2,?m) \ r" - using Trans trans_def[of r] by blast - hence "(b1,b2) \ (rel.under r ?m) \ (rel.under r ?m)" unfolding rel.under_def by simp} - thus "rel.underS ?r' (a1,a2) \ (rel.under r ?m) \ (rel.under r ?m)" by auto - qed - moreover have "wo_rel.ofilter r (rel.under r ?m)" - using Well by (simp add: wo_rel.under_ofilter) - moreover have "rel.under r ?m < Field r" - using NE rel.under_Field[of r ?m] by blast - ultimately show ?thesis by blast -qed - - -end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Fun_More.thy --- a/src/HOL/Cardinals/Fun_More.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/Fun_More.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* More on Injections, Bijections and Inverses *} theory Fun_More -imports Fun_More_LFP +imports Fun_More_FP begin diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Fun_More_FP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Cardinals/Fun_More_FP.thy Mon Nov 18 18:04:45 2013 +0100 @@ -0,0 +1,239 @@ +(* Title: HOL/Cardinals/Fun_More_FP.thy + Author: Andrei Popescu, TU Muenchen + Copyright 2012 + +More on injections, bijections and inverses (FP). +*) + +header {* More on Injections, Bijections and Inverses (FP) *} + +theory Fun_More_FP +imports "~~/src/HOL/Library/Infinite_Set" +begin + + +text {* This section proves more facts (additional to those in @{text "Fun.thy"}, +@{text "Hilbert_Choice.thy"}, @{text "Finite_Set.thy"} and @{text "Infinite_Set.thy"}), +mainly concerning injections, bijections, inverses and (numeric) cardinals of +finite sets. *} + + +subsection {* Purely functional properties *} + + +(*2*)lemma bij_betw_id_iff: +"(A = B) = (bij_betw id A B)" +by(simp add: bij_betw_def) + + +(*2*)lemma bij_betw_byWitness: +assumes LEFT: "\a \ A. f'(f a) = a" and + RIGHT: "\a' \ A'. f(f' a') = a'" and + IM1: "f ` A \ A'" and IM2: "f' ` A' \ A" +shows "bij_betw f A A'" +using assms +proof(unfold bij_betw_def inj_on_def, safe) + fix a b assume *: "a \ A" "b \ A" and **: "f a = f b" + have "a = f'(f a) \ b = f'(f b)" using * LEFT by simp + with ** show "a = b" by simp +next + fix a' assume *: "a' \ A'" + hence "f' a' \ A" using IM2 by blast + moreover + have "a' = f(f' a')" using * RIGHT by simp + ultimately show "a' \ f ` A" by blast +qed + + +(*3*)corollary notIn_Un_bij_betw: +assumes NIN: "b \ A" and NIN': "f b \ A'" and + BIJ: "bij_betw f A A'" +shows "bij_betw f (A \ {b}) (A' \ {f b})" +proof- + have "bij_betw f {b} {f b}" + unfolding bij_betw_def inj_on_def by simp + with assms show ?thesis + using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast +qed + + +(*1*)lemma notIn_Un_bij_betw3: +assumes NIN: "b \ A" and NIN': "f b \ A'" +shows "bij_betw f A A' = bij_betw f (A \ {b}) (A' \ {f b})" +proof + assume "bij_betw f A A'" + thus "bij_betw f (A \ {b}) (A' \ {f b})" + using assms notIn_Un_bij_betw[of b A f A'] by blast +next + assume *: "bij_betw f (A \ {b}) (A' \ {f b})" + have "f ` A = A'" + proof(auto) + fix a assume **: "a \ A" + hence "f a \ A' \ {f b}" using * unfolding bij_betw_def by blast + moreover + {assume "f a = f b" + hence "a = b" using * ** unfolding bij_betw_def inj_on_def by blast + with NIN ** have False by blast + } + ultimately show "f a \ A'" by blast + next + fix a' assume **: "a' \ A'" + hence "a' \ f`(A \ {b})" + using * by (auto simp add: bij_betw_def) + then obtain a where 1: "a \ A \ {b} \ f a = a'" by blast + moreover + {assume "a = b" with 1 ** NIN' have False by blast + } + ultimately have "a \ A" by blast + with 1 show "a' \ f ` A" by blast + qed + thus "bij_betw f A A'" using * bij_betw_subset[of f "A \ {b}" _ A] by blast +qed + + +subsection {* Properties involving finite and infinite sets *} + + +(*3*)lemma inj_on_finite: +assumes "inj_on f A" "f ` A \ B" "finite B" +shows "finite A" +using assms infinite_super by (fast dest: finite_imageD) + + +(*3*)lemma infinite_imp_bij_betw: +assumes INF: "infinite A" +shows "\h. bij_betw h A (A - {a})" +proof(cases "a \ A") + assume Case1: "a \ A" hence "A - {a} = A" by blast + thus ?thesis using bij_betw_id[of A] by auto +next + assume Case2: "a \ A" + have "infinite (A - {a})" using INF infinite_remove by auto + with infinite_iff_countable_subset[of "A - {a}"] obtain f::"nat \ 'a" + where 1: "inj f" and 2: "f ` UNIV \ A - {a}" by blast + obtain g where g_def: "g = (\ n. if n = 0 then a else f (Suc n))" by blast + obtain A' where A'_def: "A' = g ` UNIV" by blast + have temp: "\y. f y \ a" using 2 by blast + have 3: "inj_on g UNIV \ g ` UNIV \ A \ a \ g ` UNIV" + proof(auto simp add: Case2 g_def, unfold inj_on_def, intro ballI impI, + case_tac "x = 0", auto simp add: 2) + fix y assume "a = (if y = 0 then a else f (Suc y))" + thus "y = 0" using temp by (case_tac "y = 0", auto) + next + fix x y + assume "f (Suc x) = (if y = 0 then a else f (Suc y))" + thus "x = y" using 1 temp unfolding inj_on_def by (case_tac "y = 0", auto) + next + fix n show "f (Suc n) \ A" using 2 by blast + qed + hence 4: "bij_betw g UNIV A' \ a \ A' \ A' \ A" + using inj_on_imp_bij_betw[of g] unfolding A'_def by auto + hence 5: "bij_betw (inv g) A' UNIV" + by (auto simp add: bij_betw_inv_into) + (* *) + obtain n where "g n = a" using 3 by auto + hence 6: "bij_betw g (UNIV - {n}) (A' - {a})" + using 3 4 unfolding A'_def + by clarify (rule bij_betw_subset, auto simp: image_set_diff) + (* *) + obtain v where v_def: "v = (\ m. if m < n then m else Suc m)" by blast + have 7: "bij_betw v UNIV (UNIV - {n})" + proof(unfold bij_betw_def inj_on_def, intro conjI, clarify) + fix m1 m2 assume "v m1 = v m2" + thus "m1 = m2" + by(case_tac "m1 < n", case_tac "m2 < n", + auto simp add: inj_on_def v_def, case_tac "m2 < n", auto) + next + show "v ` UNIV = UNIV - {n}" + proof(auto simp add: v_def) + fix m assume *: "m \ n" and **: "m \ Suc ` {m'. \ m' < n}" + {assume "n \ m" with * have 71: "Suc n \ m" by auto + then obtain m' where 72: "m = Suc m'" using Suc_le_D by auto + with 71 have "n \ m'" by auto + with 72 ** have False by auto + } + thus "m < n" by force + qed + qed + (* *) + obtain h' where h'_def: "h' = g o v o (inv g)" by blast + hence 8: "bij_betw h' A' (A' - {a})" using 5 7 6 + by (auto simp add: bij_betw_trans) + (* *) + obtain h where h_def: "h = (\ b. if b \ A' then h' b else b)" by blast + have "\b \ A'. h b = h' b" unfolding h_def by auto + hence "bij_betw h A' (A' - {a})" using 8 bij_betw_cong[of A' h] by auto + moreover + {have "\b \ A - A'. h b = b" unfolding h_def by auto + hence "bij_betw h (A - A') (A - A')" + using bij_betw_cong[of "A - A'" h id] bij_betw_id[of "A - A'"] by auto + } + moreover + have "(A' Int (A - A') = {} \ A' \ (A - A') = A) \ + ((A' - {a}) Int (A - A') = {} \ (A' - {a}) \ (A - A') = A - {a})" + using 4 by blast + ultimately have "bij_betw h A (A - {a})" + using bij_betw_combine[of h A' "A' - {a}" "A - A'" "A - A'"] by simp + thus ?thesis by blast +qed + + +(*3*)lemma infinite_imp_bij_betw2: +assumes INF: "infinite A" +shows "\h. bij_betw h A (A \ {a})" +proof(cases "a \ A") + assume Case1: "a \ A" hence "A \ {a} = A" by blast + thus ?thesis using bij_betw_id[of A] by auto +next + let ?A' = "A \ {a}" + assume Case2: "a \ A" hence "A = ?A' - {a}" by blast + moreover have "infinite ?A'" using INF by auto + ultimately obtain f where "bij_betw f ?A' A" + using infinite_imp_bij_betw[of ?A' a] by auto + hence "bij_betw(inv_into ?A' f) A ?A'" using bij_betw_inv_into by blast + thus ?thesis by auto +qed + + +subsection {* Properties involving Hilbert choice *} + + +(*2*)lemma bij_betw_inv_into_left: +assumes BIJ: "bij_betw f A A'" and IN: "a \ A" +shows "(inv_into A f) (f a) = a" +using assms unfolding bij_betw_def +by clarify (rule inv_into_f_f) + +(*2*)lemma bij_betw_inv_into_right: +assumes "bij_betw f A A'" "a' \ A'" +shows "f(inv_into A f a') = a'" +using assms unfolding bij_betw_def using f_inv_into_f by force + + +(*1*)lemma bij_betw_inv_into_subset: +assumes BIJ: "bij_betw f A A'" and + SUB: "B \ A" and IM: "f ` B = B'" +shows "bij_betw (inv_into A f) B' B" +using assms unfolding bij_betw_def +by (auto intro: inj_on_inv_into) + + +subsection {* Other facts *} + + +(*2*)lemma atLeastLessThan_less_eq: +"({0.. {0.. n)" +unfolding ivl_subset by arith + + +(*2*)lemma atLeastLessThan_less_eq2: +assumes "inj_on f {0..<(m::nat)} \ f ` {0.. {0.. n" +using assms +using finite_atLeastLessThan[of m] finite_atLeastLessThan[of n] + card_atLeastLessThan[of m] card_atLeastLessThan[of n] + card_inj_on_le[of f "{0 ..< m}" "{0 ..< n}"] by auto + + + +end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Fun_More_LFP.thy --- a/src/HOL/Cardinals/Fun_More_LFP.thy Mon Nov 18 18:04:45 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,239 +0,0 @@ -(* Title: HOL/Cardinals/Fun_More_LFP.thy - Author: Andrei Popescu, TU Muenchen - Copyright 2012 - -More on injections, bijections and inverses (LFP). -*) - -header {* More on Injections, Bijections and Inverses (LFP) *} - -theory Fun_More_LFP -imports "~~/src/HOL/Library/Infinite_Set" -begin - - -text {* This section proves more facts (additional to those in @{text "Fun.thy"}, -@{text "Hilbert_Choice.thy"}, @{text "Finite_Set.thy"} and @{text "Infinite_Set.thy"}), -mainly concerning injections, bijections, inverses and (numeric) cardinals of -finite sets. *} - - -subsection {* Purely functional properties *} - - -(*2*)lemma bij_betw_id_iff: -"(A = B) = (bij_betw id A B)" -by(simp add: bij_betw_def) - - -(*2*)lemma bij_betw_byWitness: -assumes LEFT: "\a \ A. f'(f a) = a" and - RIGHT: "\a' \ A'. f(f' a') = a'" and - IM1: "f ` A \ A'" and IM2: "f' ` A' \ A" -shows "bij_betw f A A'" -using assms -proof(unfold bij_betw_def inj_on_def, safe) - fix a b assume *: "a \ A" "b \ A" and **: "f a = f b" - have "a = f'(f a) \ b = f'(f b)" using * LEFT by simp - with ** show "a = b" by simp -next - fix a' assume *: "a' \ A'" - hence "f' a' \ A" using IM2 by blast - moreover - have "a' = f(f' a')" using * RIGHT by simp - ultimately show "a' \ f ` A" by blast -qed - - -(*3*)corollary notIn_Un_bij_betw: -assumes NIN: "b \ A" and NIN': "f b \ A'" and - BIJ: "bij_betw f A A'" -shows "bij_betw f (A \ {b}) (A' \ {f b})" -proof- - have "bij_betw f {b} {f b}" - unfolding bij_betw_def inj_on_def by simp - with assms show ?thesis - using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast -qed - - -(*1*)lemma notIn_Un_bij_betw3: -assumes NIN: "b \ A" and NIN': "f b \ A'" -shows "bij_betw f A A' = bij_betw f (A \ {b}) (A' \ {f b})" -proof - assume "bij_betw f A A'" - thus "bij_betw f (A \ {b}) (A' \ {f b})" - using assms notIn_Un_bij_betw[of b A f A'] by blast -next - assume *: "bij_betw f (A \ {b}) (A' \ {f b})" - have "f ` A = A'" - proof(auto) - fix a assume **: "a \ A" - hence "f a \ A' \ {f b}" using * unfolding bij_betw_def by blast - moreover - {assume "f a = f b" - hence "a = b" using * ** unfolding bij_betw_def inj_on_def by blast - with NIN ** have False by blast - } - ultimately show "f a \ A'" by blast - next - fix a' assume **: "a' \ A'" - hence "a' \ f`(A \ {b})" - using * by (auto simp add: bij_betw_def) - then obtain a where 1: "a \ A \ {b} \ f a = a'" by blast - moreover - {assume "a = b" with 1 ** NIN' have False by blast - } - ultimately have "a \ A" by blast - with 1 show "a' \ f ` A" by blast - qed - thus "bij_betw f A A'" using * bij_betw_subset[of f "A \ {b}" _ A] by blast -qed - - -subsection {* Properties involving finite and infinite sets *} - - -(*3*)lemma inj_on_finite: -assumes "inj_on f A" "f ` A \ B" "finite B" -shows "finite A" -using assms infinite_super by (fast dest: finite_imageD) - - -(*3*)lemma infinite_imp_bij_betw: -assumes INF: "infinite A" -shows "\h. bij_betw h A (A - {a})" -proof(cases "a \ A") - assume Case1: "a \ A" hence "A - {a} = A" by blast - thus ?thesis using bij_betw_id[of A] by auto -next - assume Case2: "a \ A" - have "infinite (A - {a})" using INF infinite_remove by auto - with infinite_iff_countable_subset[of "A - {a}"] obtain f::"nat \ 'a" - where 1: "inj f" and 2: "f ` UNIV \ A - {a}" by blast - obtain g where g_def: "g = (\ n. if n = 0 then a else f (Suc n))" by blast - obtain A' where A'_def: "A' = g ` UNIV" by blast - have temp: "\y. f y \ a" using 2 by blast - have 3: "inj_on g UNIV \ g ` UNIV \ A \ a \ g ` UNIV" - proof(auto simp add: Case2 g_def, unfold inj_on_def, intro ballI impI, - case_tac "x = 0", auto simp add: 2) - fix y assume "a = (if y = 0 then a else f (Suc y))" - thus "y = 0" using temp by (case_tac "y = 0", auto) - next - fix x y - assume "f (Suc x) = (if y = 0 then a else f (Suc y))" - thus "x = y" using 1 temp unfolding inj_on_def by (case_tac "y = 0", auto) - next - fix n show "f (Suc n) \ A" using 2 by blast - qed - hence 4: "bij_betw g UNIV A' \ a \ A' \ A' \ A" - using inj_on_imp_bij_betw[of g] unfolding A'_def by auto - hence 5: "bij_betw (inv g) A' UNIV" - by (auto simp add: bij_betw_inv_into) - (* *) - obtain n where "g n = a" using 3 by auto - hence 6: "bij_betw g (UNIV - {n}) (A' - {a})" - using 3 4 unfolding A'_def - by clarify (rule bij_betw_subset, auto simp: image_set_diff) - (* *) - obtain v where v_def: "v = (\ m. if m < n then m else Suc m)" by blast - have 7: "bij_betw v UNIV (UNIV - {n})" - proof(unfold bij_betw_def inj_on_def, intro conjI, clarify) - fix m1 m2 assume "v m1 = v m2" - thus "m1 = m2" - by(case_tac "m1 < n", case_tac "m2 < n", - auto simp add: inj_on_def v_def, case_tac "m2 < n", auto) - next - show "v ` UNIV = UNIV - {n}" - proof(auto simp add: v_def) - fix m assume *: "m \ n" and **: "m \ Suc ` {m'. \ m' < n}" - {assume "n \ m" with * have 71: "Suc n \ m" by auto - then obtain m' where 72: "m = Suc m'" using Suc_le_D by auto - with 71 have "n \ m'" by auto - with 72 ** have False by auto - } - thus "m < n" by force - qed - qed - (* *) - obtain h' where h'_def: "h' = g o v o (inv g)" by blast - hence 8: "bij_betw h' A' (A' - {a})" using 5 7 6 - by (auto simp add: bij_betw_trans) - (* *) - obtain h where h_def: "h = (\ b. if b \ A' then h' b else b)" by blast - have "\b \ A'. h b = h' b" unfolding h_def by auto - hence "bij_betw h A' (A' - {a})" using 8 bij_betw_cong[of A' h] by auto - moreover - {have "\b \ A - A'. h b = b" unfolding h_def by auto - hence "bij_betw h (A - A') (A - A')" - using bij_betw_cong[of "A - A'" h id] bij_betw_id[of "A - A'"] by auto - } - moreover - have "(A' Int (A - A') = {} \ A' \ (A - A') = A) \ - ((A' - {a}) Int (A - A') = {} \ (A' - {a}) \ (A - A') = A - {a})" - using 4 by blast - ultimately have "bij_betw h A (A - {a})" - using bij_betw_combine[of h A' "A' - {a}" "A - A'" "A - A'"] by simp - thus ?thesis by blast -qed - - -(*3*)lemma infinite_imp_bij_betw2: -assumes INF: "infinite A" -shows "\h. bij_betw h A (A \ {a})" -proof(cases "a \ A") - assume Case1: "a \ A" hence "A \ {a} = A" by blast - thus ?thesis using bij_betw_id[of A] by auto -next - let ?A' = "A \ {a}" - assume Case2: "a \ A" hence "A = ?A' - {a}" by blast - moreover have "infinite ?A'" using INF by auto - ultimately obtain f where "bij_betw f ?A' A" - using infinite_imp_bij_betw[of ?A' a] by auto - hence "bij_betw(inv_into ?A' f) A ?A'" using bij_betw_inv_into by blast - thus ?thesis by auto -qed - - -subsection {* Properties involving Hilbert choice *} - - -(*2*)lemma bij_betw_inv_into_left: -assumes BIJ: "bij_betw f A A'" and IN: "a \ A" -shows "(inv_into A f) (f a) = a" -using assms unfolding bij_betw_def -by clarify (rule inv_into_f_f) - -(*2*)lemma bij_betw_inv_into_right: -assumes "bij_betw f A A'" "a' \ A'" -shows "f(inv_into A f a') = a'" -using assms unfolding bij_betw_def using f_inv_into_f by force - - -(*1*)lemma bij_betw_inv_into_subset: -assumes BIJ: "bij_betw f A A'" and - SUB: "B \ A" and IM: "f ` B = B'" -shows "bij_betw (inv_into A f) B' B" -using assms unfolding bij_betw_def -by (auto intro: inj_on_inv_into) - - -subsection {* Other facts *} - - -(*2*)lemma atLeastLessThan_less_eq: -"({0.. {0.. n)" -unfolding ivl_subset by arith - - -(*2*)lemma atLeastLessThan_less_eq2: -assumes "inj_on f {0..<(m::nat)} \ f ` {0.. {0.. n" -using assms -using finite_atLeastLessThan[of m] finite_atLeastLessThan[of n] - card_atLeastLessThan[of m] card_atLeastLessThan[of n] - card_inj_on_le[of f "{0 ..< m}" "{0 ..< n}"] by auto - - - -end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Order_Relation_More.thy --- a/src/HOL/Cardinals/Order_Relation_More.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/Order_Relation_More.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* Basics on Order-Like Relations *} theory Order_Relation_More -imports Order_Relation_More_LFP +imports Order_Relation_More_FP begin diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Order_Relation_More_FP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Cardinals/Order_Relation_More_FP.thy Mon Nov 18 18:04:45 2013 +0100 @@ -0,0 +1,230 @@ +(* Title: HOL/Cardinals/Order_Relation_More_FP.thy + Author: Andrei Popescu, TU Muenchen + Copyright 2012 + +Basics on order-like relations (FP). +*) + +header {* Basics on Order-Like Relations (FP) *} + +theory Order_Relation_More_FP +imports "~~/src/HOL/Library/Order_Relation" +begin + + +text{* In this section, we develop basic concepts and results pertaining +to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or +total relations. The development is placed on top of the definitions +from the theory @{text "Order_Relation"}. We also +further define upper and lower bounds operators. *} + + +locale rel = fixes r :: "'a rel" + +text{* The following context encompasses all this section, except +for its last subsection. In other words, for the rest of this section except its last +subsection, we consider a fixed relation @{text "r"}. *} + +context rel +begin + + +subsection {* Auxiliaries *} + + +lemma refl_on_domain: +"\refl_on A r; (a,b) : r\ \ a \ A \ b \ A" +by(auto simp add: refl_on_def) + + +corollary well_order_on_domain: +"\well_order_on A r; (a,b) \ r\ \ a \ A \ b \ A" +by(simp add: refl_on_domain order_on_defs) + + +lemma well_order_on_Field: +"well_order_on A r \ A = Field r" +by(auto simp add: refl_on_def Field_def order_on_defs) + + +lemma well_order_on_Well_order: +"well_order_on A r \ A = Field r \ Well_order r" +using well_order_on_Field by simp + + +lemma Total_subset_Id: +assumes TOT: "Total r" and SUB: "r \ Id" +shows "r = {} \ (\a. r = {(a,a)})" +proof- + {assume "r \ {}" + then obtain a b where 1: "(a,b) \ r" by fast + hence "a = b" using SUB by blast + hence 2: "(a,a) \ r" using 1 by simp + {fix c d assume "(c,d) \ r" + hence "{a,c,d} \ Field r" using 1 unfolding Field_def by blast + hence "((a,c) \ r \ (c,a) \ r \ a = c) \ + ((a,d) \ r \ (d,a) \ r \ a = d)" + using TOT unfolding total_on_def by blast + hence "a = c \ a = d" using SUB by blast + } + hence "r \ {(a,a)}" by auto + with 2 have "\a. r = {(a,a)}" by blast + } + thus ?thesis by blast +qed + + +lemma Linear_order_in_diff_Id: +assumes LI: "Linear_order r" and + IN1: "a \ Field r" and IN2: "b \ Field r" +shows "((a,b) \ r) = ((b,a) \ r - Id)" +using assms unfolding order_on_defs total_on_def antisym_def Id_def refl_on_def by force + + +subsection {* The upper and lower bounds operators *} + + +text{* Here we define upper (``above") and lower (``below") bounds operators. +We think of @{text "r"} as a {\em non-strict} relation. The suffix ``S" +at the names of some operators indicates that the bounds are strict -- e.g., +@{text "underS a"} is the set of all strict lower bounds of @{text "a"} (w.r.t. @{text "r"}). +Capitalization of the first letter in the name reminds that the operator acts on sets, rather +than on individual elements. *} + +definition under::"'a \ 'a set" +where "under a \ {b. (b,a) \ r}" + +definition underS::"'a \ 'a set" +where "underS a \ {b. b \ a \ (b,a) \ r}" + +definition Under::"'a set \ 'a set" +where "Under A \ {b \ Field r. \a \ A. (b,a) \ r}" + +definition UnderS::"'a set \ 'a set" +where "UnderS A \ {b \ Field r. \a \ A. b \ a \ (b,a) \ r}" + +definition above::"'a \ 'a set" +where "above a \ {b. (a,b) \ r}" + +definition aboveS::"'a \ 'a set" +where "aboveS a \ {b. b \ a \ (a,b) \ r}" + +definition Above::"'a set \ 'a set" +where "Above A \ {b \ Field r. \a \ A. (a,b) \ r}" + +definition AboveS::"'a set \ 'a set" +where "AboveS A \ {b \ Field r. \a \ A. b \ a \ (a,b) \ r}" +(* *) + +text{* Note: In the definitions of @{text "Above[S]"} and @{text "Under[S]"}, + we bounded comprehension by @{text "Field r"} in order to properly cover + the case of @{text "A"} being empty. *} + + +lemma underS_subset_under: "underS a \ under a" +by(auto simp add: underS_def under_def) + + +lemma underS_notIn: "a \ underS a" +by(simp add: underS_def) + + +lemma Refl_under_in: "\Refl r; a \ Field r\ \ a \ under a" +by(simp add: refl_on_def under_def) + + +lemma AboveS_disjoint: "A Int (AboveS A) = {}" +by(auto simp add: AboveS_def) + + +lemma in_AboveS_underS: "a \ Field r \ a \ AboveS (underS a)" +by(auto simp add: AboveS_def underS_def) + + +lemma Refl_under_underS: +assumes "Refl r" "a \ Field r" +shows "under a = underS a \ {a}" +unfolding under_def underS_def +using assms refl_on_def[of _ r] by fastforce + + +lemma underS_empty: "a \ Field r \ underS a = {}" +by (auto simp: Field_def underS_def) + + +lemma under_Field: "under a \ Field r" +by(unfold under_def Field_def, auto) + + +lemma underS_Field: "underS a \ Field r" +by(unfold underS_def Field_def, auto) + + +lemma underS_Field2: +"a \ Field r \ underS a < Field r" +using assms underS_notIn underS_Field by blast + + +lemma underS_Field3: +"Field r \ {} \ underS a < Field r" +by(cases "a \ Field r", simp add: underS_Field2, auto simp add: underS_empty) + + +lemma AboveS_Field: "AboveS A \ Field r" +by(unfold AboveS_def Field_def, auto) + + +lemma under_incr: +assumes TRANS: "trans r" and REL: "(a,b) \ r" +shows "under a \ under b" +proof(unfold under_def, auto) + fix x assume "(x,a) \ r" + with REL TRANS trans_def[of r] + show "(x,b) \ r" by blast +qed + + +lemma underS_incr: +assumes TRANS: "trans r" and ANTISYM: "antisym r" and + REL: "(a,b) \ r" +shows "underS a \ underS b" +proof(unfold underS_def, auto) + assume *: "b \ a" and **: "(b,a) \ r" + with ANTISYM antisym_def[of r] REL + show False by blast +next + fix x assume "x \ a" "(x,a) \ r" + with REL TRANS trans_def[of r] + show "(x,b) \ r" by blast +qed + + +lemma underS_incl_iff: +assumes LO: "Linear_order r" and + INa: "a \ Field r" and INb: "b \ Field r" +shows "(underS a \ underS b) = ((a,b) \ r)" +proof + assume "(a,b) \ r" + thus "underS a \ underS b" using LO + by (simp add: order_on_defs underS_incr) +next + assume *: "underS a \ underS b" + {assume "a = b" + hence "(a,b) \ r" using assms + by (simp add: order_on_defs refl_on_def) + } + moreover + {assume "a \ b \ (b,a) \ r" + hence "b \ underS a" unfolding underS_def by blast + hence "b \ underS b" using * by blast + hence False by (simp add: underS_notIn) + } + ultimately + show "(a,b) \ r" using assms + order_on_defs[of "Field r" r] total_on_def[of "Field r" r] by blast +qed + + +end (* context rel *) + +end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Order_Relation_More_LFP.thy --- a/src/HOL/Cardinals/Order_Relation_More_LFP.thy Mon Nov 18 18:04:45 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,230 +0,0 @@ -(* Title: HOL/Cardinals/Order_Relation_More_LFP.thy - Author: Andrei Popescu, TU Muenchen - Copyright 2012 - -Basics on order-like relations (LFP). -*) - -header {* Basics on Order-Like Relations (LFP) *} - -theory Order_Relation_More_LFP -imports "~~/src/HOL/Library/Order_Relation" -begin - - -text{* In this section, we develop basic concepts and results pertaining -to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or -total relations. The development is placed on top of the definitions -from the theory @{text "Order_Relation"}. We also -further define upper and lower bounds operators. *} - - -locale rel = fixes r :: "'a rel" - -text{* The following context encompasses all this section, except -for its last subsection. In other words, for the rest of this section except its last -subsection, we consider a fixed relation @{text "r"}. *} - -context rel -begin - - -subsection {* Auxiliaries *} - - -lemma refl_on_domain: -"\refl_on A r; (a,b) : r\ \ a \ A \ b \ A" -by(auto simp add: refl_on_def) - - -corollary well_order_on_domain: -"\well_order_on A r; (a,b) \ r\ \ a \ A \ b \ A" -by(simp add: refl_on_domain order_on_defs) - - -lemma well_order_on_Field: -"well_order_on A r \ A = Field r" -by(auto simp add: refl_on_def Field_def order_on_defs) - - -lemma well_order_on_Well_order: -"well_order_on A r \ A = Field r \ Well_order r" -using well_order_on_Field by simp - - -lemma Total_subset_Id: -assumes TOT: "Total r" and SUB: "r \ Id" -shows "r = {} \ (\a. r = {(a,a)})" -proof- - {assume "r \ {}" - then obtain a b where 1: "(a,b) \ r" by fast - hence "a = b" using SUB by blast - hence 2: "(a,a) \ r" using 1 by simp - {fix c d assume "(c,d) \ r" - hence "{a,c,d} \ Field r" using 1 unfolding Field_def by blast - hence "((a,c) \ r \ (c,a) \ r \ a = c) \ - ((a,d) \ r \ (d,a) \ r \ a = d)" - using TOT unfolding total_on_def by blast - hence "a = c \ a = d" using SUB by blast - } - hence "r \ {(a,a)}" by auto - with 2 have "\a. r = {(a,a)}" by blast - } - thus ?thesis by blast -qed - - -lemma Linear_order_in_diff_Id: -assumes LI: "Linear_order r" and - IN1: "a \ Field r" and IN2: "b \ Field r" -shows "((a,b) \ r) = ((b,a) \ r - Id)" -using assms unfolding order_on_defs total_on_def antisym_def Id_def refl_on_def by force - - -subsection {* The upper and lower bounds operators *} - - -text{* Here we define upper (``above") and lower (``below") bounds operators. -We think of @{text "r"} as a {\em non-strict} relation. The suffix ``S" -at the names of some operators indicates that the bounds are strict -- e.g., -@{text "underS a"} is the set of all strict lower bounds of @{text "a"} (w.r.t. @{text "r"}). -Capitalization of the first letter in the name reminds that the operator acts on sets, rather -than on individual elements. *} - -definition under::"'a \ 'a set" -where "under a \ {b. (b,a) \ r}" - -definition underS::"'a \ 'a set" -where "underS a \ {b. b \ a \ (b,a) \ r}" - -definition Under::"'a set \ 'a set" -where "Under A \ {b \ Field r. \a \ A. (b,a) \ r}" - -definition UnderS::"'a set \ 'a set" -where "UnderS A \ {b \ Field r. \a \ A. b \ a \ (b,a) \ r}" - -definition above::"'a \ 'a set" -where "above a \ {b. (a,b) \ r}" - -definition aboveS::"'a \ 'a set" -where "aboveS a \ {b. b \ a \ (a,b) \ r}" - -definition Above::"'a set \ 'a set" -where "Above A \ {b \ Field r. \a \ A. (a,b) \ r}" - -definition AboveS::"'a set \ 'a set" -where "AboveS A \ {b \ Field r. \a \ A. b \ a \ (a,b) \ r}" -(* *) - -text{* Note: In the definitions of @{text "Above[S]"} and @{text "Under[S]"}, - we bounded comprehension by @{text "Field r"} in order to properly cover - the case of @{text "A"} being empty. *} - - -lemma underS_subset_under: "underS a \ under a" -by(auto simp add: underS_def under_def) - - -lemma underS_notIn: "a \ underS a" -by(simp add: underS_def) - - -lemma Refl_under_in: "\Refl r; a \ Field r\ \ a \ under a" -by(simp add: refl_on_def under_def) - - -lemma AboveS_disjoint: "A Int (AboveS A) = {}" -by(auto simp add: AboveS_def) - - -lemma in_AboveS_underS: "a \ Field r \ a \ AboveS (underS a)" -by(auto simp add: AboveS_def underS_def) - - -lemma Refl_under_underS: -assumes "Refl r" "a \ Field r" -shows "under a = underS a \ {a}" -unfolding under_def underS_def -using assms refl_on_def[of _ r] by fastforce - - -lemma underS_empty: "a \ Field r \ underS a = {}" -by (auto simp: Field_def underS_def) - - -lemma under_Field: "under a \ Field r" -by(unfold under_def Field_def, auto) - - -lemma underS_Field: "underS a \ Field r" -by(unfold underS_def Field_def, auto) - - -lemma underS_Field2: -"a \ Field r \ underS a < Field r" -using assms underS_notIn underS_Field by blast - - -lemma underS_Field3: -"Field r \ {} \ underS a < Field r" -by(cases "a \ Field r", simp add: underS_Field2, auto simp add: underS_empty) - - -lemma AboveS_Field: "AboveS A \ Field r" -by(unfold AboveS_def Field_def, auto) - - -lemma under_incr: -assumes TRANS: "trans r" and REL: "(a,b) \ r" -shows "under a \ under b" -proof(unfold under_def, auto) - fix x assume "(x,a) \ r" - with REL TRANS trans_def[of r] - show "(x,b) \ r" by blast -qed - - -lemma underS_incr: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - REL: "(a,b) \ r" -shows "underS a \ underS b" -proof(unfold underS_def, auto) - assume *: "b \ a" and **: "(b,a) \ r" - with ANTISYM antisym_def[of r] REL - show False by blast -next - fix x assume "x \ a" "(x,a) \ r" - with REL TRANS trans_def[of r] - show "(x,b) \ r" by blast -qed - - -lemma underS_incl_iff: -assumes LO: "Linear_order r" and - INa: "a \ Field r" and INb: "b \ Field r" -shows "(underS a \ underS b) = ((a,b) \ r)" -proof - assume "(a,b) \ r" - thus "underS a \ underS b" using LO - by (simp add: order_on_defs underS_incr) -next - assume *: "underS a \ underS b" - {assume "a = b" - hence "(a,b) \ r" using assms - by (simp add: order_on_defs refl_on_def) - } - moreover - {assume "a \ b \ (b,a) \ r" - hence "b \ underS a" unfolding underS_def by blast - hence "b \ underS b" using * by blast - hence False by (simp add: underS_notIn) - } - ultimately - show "(a,b) \ r" using assms - order_on_defs[of "Field r" r] total_on_def[of "Field r" r] by blast -qed - - -end (* context rel *) - -end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/README.txt --- a/src/HOL/Cardinals/README.txt Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/README.txt Mon Nov 18 18:04:45 2013 +0100 @@ -89,15 +89,16 @@ Minor technicalities and naming issues: --------------------------------------- -1. Most of the definitions and theorems are proved in files suffixed with -"_LFP" or "_GFP". Bootstrapping considerations (for the (co)datatype package) made -this division desirable. +1. Most of the definitions and theorems are proved in files suffixed with "_FP". +Bootstrapping considerations (for the (co)datatype package) made this division +desirable. -2. Even though we would have preferred to use "initial segment" instead of -"order filter", we chose the latter to avoid terminological clash with -the operator "init_seg_of" from Zorn.thy. The latter expresses a related, but different -concept -- it considers a relation, rather than a set, as initial segment of a relation. +2. Even though we would have preferred to use "initial segment" instead of +"order filter", we chose the latter to avoid terminological clash with the +operator "init_seg_of" from Zorn.thy. The latter expresses a related, but +different concept -- it considers a relation, rather than a set, as initial +segment of a relation. 3. We prefer to define the upper-bound operations under, underS, @@ -148,7 +149,7 @@ Notes for anyone who would like to enrich these theories in the future -------------------------------------------------------------------------------------- -Theory Fun_More (and Fun_More_*): +Theory Fun_More (and Fun_More_FP): - Careful: "inj" is an abbreviation for "inj_on UNIV", while "bij" is not an abreviation for "bij_betw UNIV UNIV", but a defined constant; there is no "surj_betw", but only "surj". @@ -166,7 +167,7 @@ - In subsection "Other facts": -- Does the lemma "atLeastLessThan_injective" already exist anywhere? -Theory Order_Relation_More (and Order_Relation_More_*): +Theory Order_Relation_More (and Order_Relation_More_FP): - In subsection "Auxiliaries": -- Recall the lemmas "order_on_defs", "Field_def", "Domain_def", "Range_def", "converse_def". -- Recall that "refl_on r A" forces r to not be defined outside A. @@ -181,16 +182,16 @@ abbreviation "Linear_order r ≡ linear_order_on (Field r) r" abbreviation "Well_order r ≡ well_order_on (Field r) r" -Theory Wellorder_Relation (and Wellorder_Relation_*): +Theory Wellorder_Relation (and Wellorder_Relation_FP): - In subsection "Auxiliaries": recall lemmas "order_on_defs" - In subsection "The notions of maximum, minimum, supremum, successor and order filter": Should we define all constants from "wo_rel" in "rel" instead, so that their outside definition not be conditional in "wo_rel r"? -Theory Wellfounded_More (and Wellfounded_More_*): +Theory Wellfounded_More (and Wellfounded_More_FP): Recall the lemmas "wfrec" and "wf_induct". -Theory Wellorder_Embedding (and Wellorder_Embedding_*): +Theory Wellorder_Embedding (and Wellorder_Embedding_FP): - Recall "inj_on_def" and "bij_betw_def". - Line 5 in the proof of lemma embed_in_Field: we have to figure out for this and many other situations: Why did it work without annotations to Refl_under_in? @@ -200,7 +201,7 @@ making impossible to debug theorem instantiations. - At lemma "embed_unique": If we add the attribute "rule format" at lemma, we get an error at qed. -Theory Constructions_on_Wellorders (and Constructions_on_Wellorders_*): +Theory Constructions_on_Wellorders (and Constructions_on_Wellorders_FP): - Some of the lemmas in this section are about more general kinds of relations than well-orders, but it is not clear whether they are useful in such more general contexts. - Recall that "equiv" does not have the "equiv_on" and "Equiv" versions, @@ -208,7 +209,7 @@ - The lemmas "ord_trans" are not clearly useful together, as their employment within blast or auto tends to diverge. -Theory Cardinal_Order_Relation (and Cardinal_Order_Relation_*): +Theory Cardinal_Order_Relation (and Cardinal_Order_Relation_FP): - Careful: if "|..|" meets an outer parehthesis, an extra space needs to be inserted, as in "( |A| )". - At lemma like ordLeq_Sigma_mono1: Not worth stating something like ordLeq_Sigma_mono2 -- diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellfounded_More.thy --- a/src/HOL/Cardinals/Wellfounded_More.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/Wellfounded_More.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* More on Well-Founded Relations *} theory Wellfounded_More -imports Wellfounded_More_LFP Order_Relation_More +imports Wellfounded_More_FP Order_Relation_More begin diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellfounded_More_FP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Cardinals/Wellfounded_More_FP.thy Mon Nov 18 18:04:45 2013 +0100 @@ -0,0 +1,194 @@ +(* Title: HOL/Cardinals/Wellfounded_More_FP.thy + Author: Andrei Popescu, TU Muenchen + Copyright 2012 + +More on well-founded relations (FP). +*) + +header {* More on Well-Founded Relations (FP) *} + +theory Wellfounded_More_FP +imports Order_Relation_More_FP "~~/src/HOL/Library/Wfrec" +begin + + +text {* This section contains some variations of results in the theory +@{text "Wellfounded.thy"}: +\begin{itemize} +\item means for slightly more direct definitions by well-founded recursion; +\item variations of well-founded induction; +\item means for proving a linear order to be a well-order. +\end{itemize} *} + + +subsection {* Well-founded recursion via genuine fixpoints *} + + +(*2*)lemma wfrec_fixpoint: +fixes r :: "('a * 'a) set" and + H :: "('a \ 'b) \ 'a \ 'b" +assumes WF: "wf r" and ADM: "adm_wf r H" +shows "wfrec r H = H (wfrec r H)" +proof(rule ext) + fix x + have "wfrec r H x = H (cut (wfrec r H) r x) x" + using wfrec[of r H] WF by simp + also + {have "\ y. (y,x) : r \ (cut (wfrec r H) r x) y = (wfrec r H) y" + by (auto simp add: cut_apply) + hence "H (cut (wfrec r H) r x) x = H (wfrec r H) x" + using ADM adm_wf_def[of r H] by auto + } + finally show "wfrec r H x = H (wfrec r H) x" . +qed + + + +subsection {* Characterizations of well-founded-ness *} + + +text {* A transitive relation is well-founded iff it is ``locally" well-founded, +i.e., iff its restriction to the lower bounds of of any element is well-founded. *} + +(*3*)lemma trans_wf_iff: +assumes "trans r" +shows "wf r = (\a. wf(r Int (r^-1``{a} \ r^-1``{a})))" +proof- + obtain R where R_def: "R = (\ a. r Int (r^-1``{a} \ r^-1``{a}))" by blast + {assume *: "wf r" + {fix a + have "wf(R a)" + using * R_def wf_subset[of r "R a"] by auto + } + } + (* *) + moreover + {assume *: "\a. wf(R a)" + have "wf r" + proof(unfold wf_def, clarify) + fix phi a + assume **: "\a. (\b. (b,a) \ r \ phi b) \ phi a" + obtain chi where chi_def: "chi = (\b. (b,a) \ r \ phi b)" by blast + with * have "wf (R a)" by auto + hence "(\b. (\c. (c,b) \ R a \ chi c) \ chi b) \ (\b. chi b)" + unfolding wf_def by blast + moreover + have "\b. (\c. (c,b) \ R a \ chi c) \ chi b" + proof(auto simp add: chi_def R_def) + fix b + assume 1: "(b,a) \ r" and 2: "\c. (c, b) \ r \ (c, a) \ r \ phi c" + hence "\c. (c, b) \ r \ phi c" + using assms trans_def[of r] by blast + thus "phi b" using ** by blast + qed + ultimately have "\b. chi b" by (rule mp) + with ** chi_def show "phi a" by blast + qed + } + ultimately show ?thesis using R_def by blast +qed + + +text {* The next lemma is a variation of @{text "wf_eq_minimal"} from Wellfounded, +allowing one to assume the set included in the field. *} + +(*2*)lemma wf_eq_minimal2: +"wf r = (\A. A <= Field r \ A \ {} \ (\a \ A. \a' \ A. \ (a',a) \ r))" +proof- + let ?phi = "\ A. A \ {} \ (\a \ A. \a' \ A. \ (a',a) \ r)" + have "wf r = (\A. ?phi A)" + by (auto simp: ex_in_conv [THEN sym], erule wfE_min, assumption, blast) + (rule wfI_min, metis) + (* *) + also have "(\A. ?phi A) = (\B \ Field r. ?phi B)" + proof + assume "\A. ?phi A" + thus "\B \ Field r. ?phi B" by simp + next + assume *: "\B \ Field r. ?phi B" + show "\A. ?phi A" + proof(clarify) + fix A::"'a set" assume **: "A \ {}" + obtain B where B_def: "B = A Int (Field r)" by blast + show "\a \ A. \a' \ A. (a',a) \ r" + proof(cases "B = {}") + assume Case1: "B = {}" + obtain a where 1: "a \ A \ a \ Field r" + using ** Case1 unfolding B_def by blast + hence "\a' \ A. (a',a) \ r" using 1 unfolding Field_def by blast + thus ?thesis using 1 by blast + next + assume Case2: "B \ {}" have 1: "B \ Field r" unfolding B_def by blast + obtain a where 2: "a \ B \ (\a' \ B. (a',a) \ r)" + using Case2 1 * by blast + have "\a' \ A. (a',a) \ r" + proof(clarify) + fix a' assume "a' \ A" and **: "(a',a) \ r" + hence "a' \ B" unfolding B_def Field_def by blast + thus False using 2 ** by blast + qed + thus ?thesis using 2 unfolding B_def by blast + qed + qed + qed + finally show ?thesis by blast +qed + +subsection {* Characterizations of well-founded-ness *} + +text {* The next lemma and its corollary enable one to prove that +a linear order is a well-order in a way which is more standard than +via well-founded-ness of the strict version of the relation. *} + +(*3*) +lemma Linear_order_wf_diff_Id: +assumes LI: "Linear_order r" +shows "wf(r - Id) = (\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r))" +proof(cases "r \ Id") + assume Case1: "r \ Id" + hence temp: "r - Id = {}" by blast + hence "wf(r - Id)" by (simp add: temp) + moreover + {fix A assume *: "A \ Field r" and **: "A \ {}" + obtain a where 1: "r = {} \ r = {(a,a)}" using LI + unfolding order_on_defs using Case1 rel.Total_subset_Id by metis + hence "A = {a} \ r = {(a,a)}" using * ** unfolding Field_def by blast + hence "\a \ A. \a' \ A. (a,a') \ r" using 1 by blast + } + ultimately show ?thesis by blast +next + assume Case2: "\ r \ Id" + hence 1: "Field r = Field(r - Id)" using Total_Id_Field LI + unfolding order_on_defs by blast + show ?thesis + proof + assume *: "wf(r - Id)" + show "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r)" + proof(clarify) + fix A assume **: "A \ Field r" and ***: "A \ {}" + hence "\a \ A. \a' \ A. (a',a) \ r - Id" + using 1 * unfolding wf_eq_minimal2 by simp + moreover have "\a \ A. \a' \ A. ((a,a') \ r) = ((a',a) \ r - Id)" + using rel.Linear_order_in_diff_Id[of r] ** LI by blast + ultimately show "\a \ A. \a' \ A. (a,a') \ r" by blast + qed + next + assume *: "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r)" + show "wf(r - Id)" + proof(unfold wf_eq_minimal2, clarify) + fix A assume **: "A \ Field(r - Id)" and ***: "A \ {}" + hence "\a \ A. \a' \ A. (a,a') \ r" + using 1 * by simp + moreover have "\a \ A. \a' \ A. ((a,a') \ r) = ((a',a) \ r - Id)" + using rel.Linear_order_in_diff_Id[of r] ** LI mono_Field[of "r - Id" r] by blast + ultimately show "\a \ A. \a' \ A. (a',a) \ r - Id" by blast + qed + qed +qed + +(*3*)corollary Linear_order_Well_order_iff: +assumes "Linear_order r" +shows "Well_order r = (\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r))" +using assms unfolding well_order_on_def using Linear_order_wf_diff_Id[of r] by blast + +end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellfounded_More_LFP.thy --- a/src/HOL/Cardinals/Wellfounded_More_LFP.thy Mon Nov 18 18:04:45 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,194 +0,0 @@ -(* Title: HOL/Cardinals/Wellfounded_More_LFP.thy - Author: Andrei Popescu, TU Muenchen - Copyright 2012 - -More on well-founded relations (LFP). -*) - -header {* More on Well-Founded Relations (LFP) *} - -theory Wellfounded_More_LFP -imports Order_Relation_More_LFP "~~/src/HOL/Library/Wfrec" -begin - - -text {* This section contains some variations of results in the theory -@{text "Wellfounded.thy"}: -\begin{itemize} -\item means for slightly more direct definitions by well-founded recursion; -\item variations of well-founded induction; -\item means for proving a linear order to be a well-order. -\end{itemize} *} - - -subsection {* Well-founded recursion via genuine fixpoints *} - - -(*2*)lemma wfrec_fixpoint: -fixes r :: "('a * 'a) set" and - H :: "('a \ 'b) \ 'a \ 'b" -assumes WF: "wf r" and ADM: "adm_wf r H" -shows "wfrec r H = H (wfrec r H)" -proof(rule ext) - fix x - have "wfrec r H x = H (cut (wfrec r H) r x) x" - using wfrec[of r H] WF by simp - also - {have "\ y. (y,x) : r \ (cut (wfrec r H) r x) y = (wfrec r H) y" - by (auto simp add: cut_apply) - hence "H (cut (wfrec r H) r x) x = H (wfrec r H) x" - using ADM adm_wf_def[of r H] by auto - } - finally show "wfrec r H x = H (wfrec r H) x" . -qed - - - -subsection {* Characterizations of well-founded-ness *} - - -text {* A transitive relation is well-founded iff it is ``locally" well-founded, -i.e., iff its restriction to the lower bounds of of any element is well-founded. *} - -(*3*)lemma trans_wf_iff: -assumes "trans r" -shows "wf r = (\a. wf(r Int (r^-1``{a} \ r^-1``{a})))" -proof- - obtain R where R_def: "R = (\ a. r Int (r^-1``{a} \ r^-1``{a}))" by blast - {assume *: "wf r" - {fix a - have "wf(R a)" - using * R_def wf_subset[of r "R a"] by auto - } - } - (* *) - moreover - {assume *: "\a. wf(R a)" - have "wf r" - proof(unfold wf_def, clarify) - fix phi a - assume **: "\a. (\b. (b,a) \ r \ phi b) \ phi a" - obtain chi where chi_def: "chi = (\b. (b,a) \ r \ phi b)" by blast - with * have "wf (R a)" by auto - hence "(\b. (\c. (c,b) \ R a \ chi c) \ chi b) \ (\b. chi b)" - unfolding wf_def by blast - moreover - have "\b. (\c. (c,b) \ R a \ chi c) \ chi b" - proof(auto simp add: chi_def R_def) - fix b - assume 1: "(b,a) \ r" and 2: "\c. (c, b) \ r \ (c, a) \ r \ phi c" - hence "\c. (c, b) \ r \ phi c" - using assms trans_def[of r] by blast - thus "phi b" using ** by blast - qed - ultimately have "\b. chi b" by (rule mp) - with ** chi_def show "phi a" by blast - qed - } - ultimately show ?thesis using R_def by blast -qed - - -text {* The next lemma is a variation of @{text "wf_eq_minimal"} from Wellfounded, -allowing one to assume the set included in the field. *} - -(*2*)lemma wf_eq_minimal2: -"wf r = (\A. A <= Field r \ A \ {} \ (\a \ A. \a' \ A. \ (a',a) \ r))" -proof- - let ?phi = "\ A. A \ {} \ (\a \ A. \a' \ A. \ (a',a) \ r)" - have "wf r = (\A. ?phi A)" - by (auto simp: ex_in_conv [THEN sym], erule wfE_min, assumption, blast) - (rule wfI_min, metis) - (* *) - also have "(\A. ?phi A) = (\B \ Field r. ?phi B)" - proof - assume "\A. ?phi A" - thus "\B \ Field r. ?phi B" by simp - next - assume *: "\B \ Field r. ?phi B" - show "\A. ?phi A" - proof(clarify) - fix A::"'a set" assume **: "A \ {}" - obtain B where B_def: "B = A Int (Field r)" by blast - show "\a \ A. \a' \ A. (a',a) \ r" - proof(cases "B = {}") - assume Case1: "B = {}" - obtain a where 1: "a \ A \ a \ Field r" - using ** Case1 unfolding B_def by blast - hence "\a' \ A. (a',a) \ r" using 1 unfolding Field_def by blast - thus ?thesis using 1 by blast - next - assume Case2: "B \ {}" have 1: "B \ Field r" unfolding B_def by blast - obtain a where 2: "a \ B \ (\a' \ B. (a',a) \ r)" - using Case2 1 * by blast - have "\a' \ A. (a',a) \ r" - proof(clarify) - fix a' assume "a' \ A" and **: "(a',a) \ r" - hence "a' \ B" unfolding B_def Field_def by blast - thus False using 2 ** by blast - qed - thus ?thesis using 2 unfolding B_def by blast - qed - qed - qed - finally show ?thesis by blast -qed - -subsection {* Characterizations of well-founded-ness *} - -text {* The next lemma and its corollary enable one to prove that -a linear order is a well-order in a way which is more standard than -via well-founded-ness of the strict version of the relation. *} - -(*3*) -lemma Linear_order_wf_diff_Id: -assumes LI: "Linear_order r" -shows "wf(r - Id) = (\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r))" -proof(cases "r \ Id") - assume Case1: "r \ Id" - hence temp: "r - Id = {}" by blast - hence "wf(r - Id)" by (simp add: temp) - moreover - {fix A assume *: "A \ Field r" and **: "A \ {}" - obtain a where 1: "r = {} \ r = {(a,a)}" using LI - unfolding order_on_defs using Case1 rel.Total_subset_Id by metis - hence "A = {a} \ r = {(a,a)}" using * ** unfolding Field_def by blast - hence "\a \ A. \a' \ A. (a,a') \ r" using 1 by blast - } - ultimately show ?thesis by blast -next - assume Case2: "\ r \ Id" - hence 1: "Field r = Field(r - Id)" using Total_Id_Field LI - unfolding order_on_defs by blast - show ?thesis - proof - assume *: "wf(r - Id)" - show "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r)" - proof(clarify) - fix A assume **: "A \ Field r" and ***: "A \ {}" - hence "\a \ A. \a' \ A. (a',a) \ r - Id" - using 1 * unfolding wf_eq_minimal2 by simp - moreover have "\a \ A. \a' \ A. ((a,a') \ r) = ((a',a) \ r - Id)" - using rel.Linear_order_in_diff_Id[of r] ** LI by blast - ultimately show "\a \ A. \a' \ A. (a,a') \ r" by blast - qed - next - assume *: "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r)" - show "wf(r - Id)" - proof(unfold wf_eq_minimal2, clarify) - fix A assume **: "A \ Field(r - Id)" and ***: "A \ {}" - hence "\a \ A. \a' \ A. (a,a') \ r" - using 1 * by simp - moreover have "\a \ A. \a' \ A. ((a,a') \ r) = ((a',a) \ r - Id)" - using rel.Linear_order_in_diff_Id[of r] ** LI mono_Field[of "r - Id" r] by blast - ultimately show "\a \ A. \a' \ A. (a',a) \ r - Id" by blast - qed - qed -qed - -(*3*)corollary Linear_order_Well_order_iff: -assumes "Linear_order r" -shows "Well_order r = (\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r))" -using assms unfolding well_order_on_def using Linear_order_wf_diff_Id[of r] by blast - -end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellorder_Embedding.thy --- a/src/HOL/Cardinals/Wellorder_Embedding.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/Wellorder_Embedding.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* Well-Order Embeddings *} theory Wellorder_Embedding -imports Wellorder_Embedding_LFP Fun_More Wellorder_Relation +imports Wellorder_Embedding_FP Fun_More Wellorder_Relation begin diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellorder_Embedding_FP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Cardinals/Wellorder_Embedding_FP.thy Mon Nov 18 18:04:45 2013 +0100 @@ -0,0 +1,1146 @@ +(* Title: HOL/Cardinals/Wellorder_Embedding_FP.thy + Author: Andrei Popescu, TU Muenchen + Copyright 2012 + +Well-order embeddings (FP). +*) + +header {* Well-Order Embeddings (FP) *} + +theory Wellorder_Embedding_FP +imports "~~/src/HOL/Library/Zorn" Fun_More_FP Wellorder_Relation_FP +begin + + +text{* In this section, we introduce well-order {\em embeddings} and {\em isomorphisms} and +prove their basic properties. The notion of embedding is considered from the point +of view of the theory of ordinals, and therefore requires the source to be injected +as an {\em initial segment} (i.e., {\em order filter}) of the target. A main result +of this section is the existence of embeddings (in one direction or another) between +any two well-orders, having as a consequence the fact that, given any two sets on +any two types, one is smaller than (i.e., can be injected into) the other. *} + + +subsection {* Auxiliaries *} + +lemma UNION_inj_on_ofilter: +assumes WELL: "Well_order r" and + OF: "\ i. i \ I \ wo_rel.ofilter r (A i)" and + INJ: "\ i. i \ I \ inj_on f (A i)" +shows "inj_on f (\ i \ I. A i)" +proof- + have "wo_rel r" using WELL by (simp add: wo_rel_def) + hence "\ i j. \i \ I; j \ I\ \ A i <= A j \ A j <= A i" + using wo_rel.ofilter_linord[of r] OF by blast + with WELL INJ show ?thesis + by (auto simp add: inj_on_UNION_chain) +qed + + +lemma under_underS_bij_betw: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + IN: "a \ Field r" and IN': "f a \ Field r'" and + BIJ: "bij_betw f (rel.underS r a) (rel.underS r' (f a))" +shows "bij_betw f (rel.under r a) (rel.under r' (f a))" +proof- + have "a \ rel.underS r a \ f a \ rel.underS r' (f a)" + unfolding rel.underS_def by auto + moreover + {have "Refl r \ Refl r'" using WELL WELL' + by (auto simp add: order_on_defs) + hence "rel.under r a = rel.underS r a \ {a} \ + rel.under r' (f a) = rel.underS r' (f a) \ {f a}" + using IN IN' by(auto simp add: rel.Refl_under_underS) + } + ultimately show ?thesis + using BIJ notIn_Un_bij_betw[of a "rel.underS r a" f "rel.underS r' (f a)"] by auto +qed + + + +subsection {* (Well-order) embeddings, strict embeddings, isomorphisms and order-compatible +functions *} + + +text{* Standardly, a function is an embedding of a well-order in another if it injectively and +order-compatibly maps the former into an order filter of the latter. +Here we opt for a more succinct definition (operator @{text "embed"}), +asking that, for any element in the source, the function should be a bijection +between the set of strict lower bounds of that element +and the set of strict lower bounds of its image. (Later we prove equivalence with +the standard definition -- lemma @{text "embed_iff_compat_inj_on_ofilter"}.) +A {\em strict embedding} (operator @{text "embedS"}) is a non-bijective embedding +and an isomorphism (operator @{text "iso"}) is a bijective embedding. *} + + +definition embed :: "'a rel \ 'a' rel \ ('a \ 'a') \ bool" +where +"embed r r' f \ \a \ Field r. bij_betw f (rel.under r a) (rel.under r' (f a))" + + +lemmas embed_defs = embed_def embed_def[abs_def] + + +text {* Strict embeddings: *} + +definition embedS :: "'a rel \ 'a' rel \ ('a \ 'a') \ bool" +where +"embedS r r' f \ embed r r' f \ \ bij_betw f (Field r) (Field r')" + + +lemmas embedS_defs = embedS_def embedS_def[abs_def] + + +definition iso :: "'a rel \ 'a' rel \ ('a \ 'a') \ bool" +where +"iso r r' f \ embed r r' f \ bij_betw f (Field r) (Field r')" + + +lemmas iso_defs = iso_def iso_def[abs_def] + + +definition compat :: "'a rel \ 'a' rel \ ('a \ 'a') \ bool" +where +"compat r r' f \ \a b. (a,b) \ r \ (f a, f b) \ r'" + + +lemma compat_wf: +assumes CMP: "compat r r' f" and WF: "wf r'" +shows "wf r" +proof- + have "r \ inv_image r' f" + unfolding inv_image_def using CMP + by (auto simp add: compat_def) + with WF show ?thesis + using wf_inv_image[of r' f] wf_subset[of "inv_image r' f"] by auto +qed + + +lemma id_embed: "embed r r id" +by(auto simp add: id_def embed_def bij_betw_def) + + +lemma id_iso: "iso r r id" +by(auto simp add: id_def embed_def iso_def bij_betw_def) + + +lemma embed_in_Field: +assumes WELL: "Well_order r" and + EMB: "embed r r' f" and IN: "a \ Field r" +shows "f a \ Field r'" +proof- + have Well: "wo_rel r" + using WELL by (auto simp add: wo_rel_def) + hence 1: "Refl r" + by (auto simp add: wo_rel.REFL) + hence "a \ rel.under r a" using IN rel.Refl_under_in by fastforce + hence "f a \ rel.under r' (f a)" + using EMB IN by (auto simp add: embed_def bij_betw_def) + thus ?thesis unfolding Field_def + by (auto simp: rel.under_def) +qed + + +lemma comp_embed: +assumes WELL: "Well_order r" and + EMB: "embed r r' f" and EMB': "embed r' r'' f'" +shows "embed r r'' (f' o f)" +proof(unfold embed_def, auto) + fix a assume *: "a \ Field r" + hence "bij_betw f (rel.under r a) (rel.under r' (f a))" + using embed_def[of r] EMB by auto + moreover + {have "f a \ Field r'" + using EMB WELL * by (auto simp add: embed_in_Field) + hence "bij_betw f' (rel.under r' (f a)) (rel.under r'' (f' (f a)))" + using embed_def[of r'] EMB' by auto + } + ultimately + show "bij_betw (f' \ f) (rel.under r a) (rel.under r'' (f'(f a)))" + by(auto simp add: bij_betw_trans) +qed + + +lemma comp_iso: +assumes WELL: "Well_order r" and + EMB: "iso r r' f" and EMB': "iso r' r'' f'" +shows "iso r r'' (f' o f)" +using assms unfolding iso_def +by (auto simp add: comp_embed bij_betw_trans) + + +text{* That @{text "embedS"} is also preserved by function composition shall be proved only later. *} + + +lemma embed_Field: +"\Well_order r; embed r r' f\ \ f`(Field r) \ Field r'" +by (auto simp add: embed_in_Field) + + +lemma embed_preserves_ofilter: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r r' f" and OF: "wo_rel.ofilter r A" +shows "wo_rel.ofilter r' (f`A)" +proof- + (* Preliminary facts *) + from WELL have Well: "wo_rel r" unfolding wo_rel_def . + from WELL' have Well': "wo_rel r'" unfolding wo_rel_def . + from OF have 0: "A \ Field r" by(auto simp add: Well wo_rel.ofilter_def) + (* Main proof *) + show ?thesis using Well' WELL EMB 0 embed_Field[of r r' f] + proof(unfold wo_rel.ofilter_def, auto simp add: image_def) + fix a b' + assume *: "a \ A" and **: "b' \ rel.under r' (f a)" + hence "a \ Field r" using 0 by auto + hence "bij_betw f (rel.under r a) (rel.under r' (f a))" + using * EMB by (auto simp add: embed_def) + hence "f`(rel.under r a) = rel.under r' (f a)" + by (simp add: bij_betw_def) + with ** image_def[of f "rel.under r a"] obtain b where + 1: "b \ rel.under r a \ b' = f b" by blast + hence "b \ A" using Well * OF + by (auto simp add: wo_rel.ofilter_def) + with 1 show "\b \ A. b' = f b" by blast + qed +qed + + +lemma embed_Field_ofilter: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r r' f" +shows "wo_rel.ofilter r' (f`(Field r))" +proof- + have "wo_rel.ofilter r (Field r)" + using WELL by (auto simp add: wo_rel_def wo_rel.Field_ofilter) + with WELL WELL' EMB + show ?thesis by (auto simp add: embed_preserves_ofilter) +qed + + +lemma embed_compat: +assumes EMB: "embed r r' f" +shows "compat r r' f" +proof(unfold compat_def, clarify) + fix a b + assume *: "(a,b) \ r" + hence 1: "b \ Field r" using Field_def[of r] by blast + have "a \ rel.under r b" + using * rel.under_def[of r] by simp + hence "f a \ rel.under r' (f b)" + using EMB embed_def[of r r' f] + bij_betw_def[of f "rel.under r b" "rel.under r' (f b)"] + image_def[of f "rel.under r b"] 1 by auto + thus "(f a, f b) \ r'" + by (auto simp add: rel.under_def) +qed + + +lemma embed_inj_on: +assumes WELL: "Well_order r" and EMB: "embed r r' f" +shows "inj_on f (Field r)" +proof(unfold inj_on_def, clarify) + (* Preliminary facts *) + from WELL have Well: "wo_rel r" unfolding wo_rel_def . + with wo_rel.TOTAL[of r] + have Total: "Total r" by simp + from Well wo_rel.REFL[of r] + have Refl: "Refl r" by simp + (* Main proof *) + fix a b + assume *: "a \ Field r" and **: "b \ Field r" and + ***: "f a = f b" + hence 1: "a \ Field r \ b \ Field r" + unfolding Field_def by auto + {assume "(a,b) \ r" + hence "a \ rel.under r b \ b \ rel.under r b" + using Refl by(auto simp add: rel.under_def refl_on_def) + hence "a = b" + using EMB 1 *** + by (auto simp add: embed_def bij_betw_def inj_on_def) + } + moreover + {assume "(b,a) \ r" + hence "a \ rel.under r a \ b \ rel.under r a" + using Refl by(auto simp add: rel.under_def refl_on_def) + hence "a = b" + using EMB 1 *** + by (auto simp add: embed_def bij_betw_def inj_on_def) + } + ultimately + show "a = b" using Total 1 + by (auto simp add: total_on_def) +qed + + +lemma embed_underS: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r r' f" and IN: "a \ Field r" +shows "bij_betw f (rel.underS r a) (rel.underS r' (f a))" +proof- + have "bij_betw f (rel.under r a) (rel.under r' (f a))" + using assms by (auto simp add: embed_def) + moreover + {have "f a \ Field r'" using assms embed_Field[of r r' f] by auto + hence "rel.under r a = rel.underS r a \ {a} \ + rel.under r' (f a) = rel.underS r' (f a) \ {f a}" + using assms by (auto simp add: order_on_defs rel.Refl_under_underS) + } + moreover + {have "a \ rel.underS r a \ f a \ rel.underS r' (f a)" + unfolding rel.underS_def by blast + } + ultimately show ?thesis + by (auto simp add: notIn_Un_bij_betw3) +qed + + +lemma embed_iff_compat_inj_on_ofilter: +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "embed r r' f = (compat r r' f \ inj_on f (Field r) \ wo_rel.ofilter r' (f`(Field r)))" +using assms +proof(auto simp add: embed_compat embed_inj_on embed_Field_ofilter, + unfold embed_def, auto) (* get rid of one implication *) + fix a + assume *: "inj_on f (Field r)" and + **: "compat r r' f" and + ***: "wo_rel.ofilter r' (f`(Field r))" and + ****: "a \ Field r" + (* Preliminary facts *) + have Well: "wo_rel r" + using WELL wo_rel_def[of r] by simp + hence Refl: "Refl r" + using wo_rel.REFL[of r] by simp + have Total: "Total r" + using Well wo_rel.TOTAL[of r] by simp + have Well': "wo_rel r'" + using WELL' wo_rel_def[of r'] by simp + hence Antisym': "antisym r'" + using wo_rel.ANTISYM[of r'] by simp + have "(a,a) \ r" + using **** Well wo_rel.REFL[of r] + refl_on_def[of _ r] by auto + hence "(f a, f a) \ r'" + using ** by(auto simp add: compat_def) + hence 0: "f a \ Field r'" + unfolding Field_def by auto + have "f a \ f`(Field r)" + using **** by auto + hence 2: "rel.under r' (f a) \ f`(Field r)" + using Well' *** wo_rel.ofilter_def[of r' "f`(Field r)"] by fastforce + (* Main proof *) + show "bij_betw f (rel.under r a) (rel.under r' (f a))" + proof(unfold bij_betw_def, auto) + show "inj_on f (rel.under r a)" + using * + by (auto simp add: rel.under_Field subset_inj_on) + next + fix b assume "b \ rel.under r a" + thus "f b \ rel.under r' (f a)" + unfolding rel.under_def using ** + by (auto simp add: compat_def) + next + fix b' assume *****: "b' \ rel.under r' (f a)" + hence "b' \ f`(Field r)" + using 2 by auto + with Field_def[of r] obtain b where + 3: "b \ Field r" and 4: "b' = f b" by auto + have "(b,a): r" + proof- + {assume "(a,b) \ r" + with ** 4 have "(f a, b'): r'" + by (auto simp add: compat_def) + with ***** Antisym' have "f a = b'" + by(auto simp add: rel.under_def antisym_def) + with 3 **** 4 * have "a = b" + by(auto simp add: inj_on_def) + } + moreover + {assume "a = b" + hence "(b,a) \ r" using Refl **** 3 + by (auto simp add: refl_on_def) + } + ultimately + show ?thesis using Total **** 3 by (fastforce simp add: total_on_def) + qed + with 4 show "b' \ f`(rel.under r a)" + unfolding rel.under_def by auto + qed +qed + + +lemma inv_into_ofilter_embed: +assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and + BIJ: "\b \ A. bij_betw f (rel.under r b) (rel.under r' (f b))" and + IMAGE: "f ` A = Field r'" +shows "embed r' r (inv_into A f)" +proof- + (* Preliminary facts *) + have Well: "wo_rel r" + using WELL wo_rel_def[of r] by simp + have Refl: "Refl r" + using Well wo_rel.REFL[of r] by simp + have Total: "Total r" + using Well wo_rel.TOTAL[of r] by simp + (* Main proof *) + have 1: "bij_betw f A (Field r')" + proof(unfold bij_betw_def inj_on_def, auto simp add: IMAGE) + fix b1 b2 + assume *: "b1 \ A" and **: "b2 \ A" and + ***: "f b1 = f b2" + have 11: "b1 \ Field r \ b2 \ Field r" + using * ** Well OF by (auto simp add: wo_rel.ofilter_def) + moreover + {assume "(b1,b2) \ r" + hence "b1 \ rel.under r b2 \ b2 \ rel.under r b2" + unfolding rel.under_def using 11 Refl + by (auto simp add: refl_on_def) + hence "b1 = b2" using BIJ * ** *** + by (auto simp add: bij_betw_def inj_on_def) + } + moreover + {assume "(b2,b1) \ r" + hence "b1 \ rel.under r b1 \ b2 \ rel.under r b1" + unfolding rel.under_def using 11 Refl + by (auto simp add: refl_on_def) + hence "b1 = b2" using BIJ * ** *** + by (auto simp add: bij_betw_def inj_on_def) + } + ultimately + show "b1 = b2" + using Total by (auto simp add: total_on_def) + qed + (* *) + let ?f' = "(inv_into A f)" + (* *) + have 2: "\b \ A. bij_betw ?f' (rel.under r' (f b)) (rel.under r b)" + proof(clarify) + fix b assume *: "b \ A" + hence "rel.under r b \ A" + using Well OF by(auto simp add: wo_rel.ofilter_def) + moreover + have "f ` (rel.under r b) = rel.under r' (f b)" + using * BIJ by (auto simp add: bij_betw_def) + ultimately + show "bij_betw ?f' (rel.under r' (f b)) (rel.under r b)" + using 1 by (auto simp add: bij_betw_inv_into_subset) + qed + (* *) + have 3: "\b' \ Field r'. bij_betw ?f' (rel.under r' b') (rel.under r (?f' b'))" + proof(clarify) + fix b' assume *: "b' \ Field r'" + have "b' = f (?f' b')" using * 1 + by (auto simp add: bij_betw_inv_into_right) + moreover + {obtain b where 31: "b \ A" and "f b = b'" using IMAGE * by force + hence "?f' b' = b" using 1 by (auto simp add: bij_betw_inv_into_left) + with 31 have "?f' b' \ A" by auto + } + ultimately + show "bij_betw ?f' (rel.under r' b') (rel.under r (?f' b'))" + using 2 by auto + qed + (* *) + thus ?thesis unfolding embed_def . +qed + + +lemma inv_into_underS_embed: +assumes WELL: "Well_order r" and + BIJ: "\b \ rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))" and + IN: "a \ Field r" and + IMAGE: "f ` (rel.underS r a) = Field r'" +shows "embed r' r (inv_into (rel.underS r a) f)" +using assms +by(auto simp add: wo_rel_def wo_rel.underS_ofilter inv_into_ofilter_embed) + + +lemma inv_into_Field_embed: +assumes WELL: "Well_order r" and EMB: "embed r r' f" and + IMAGE: "Field r' \ f ` (Field r)" +shows "embed r' r (inv_into (Field r) f)" +proof- + have "(\b \ Field r. bij_betw f (rel.under r b) (rel.under r' (f b)))" + using EMB by (auto simp add: embed_def) + moreover + have "f ` (Field r) \ Field r'" + using EMB WELL by (auto simp add: embed_Field) + ultimately + show ?thesis using assms + by(auto simp add: wo_rel_def wo_rel.Field_ofilter inv_into_ofilter_embed) +qed + + +lemma inv_into_Field_embed_bij_betw: +assumes WELL: "Well_order r" and + EMB: "embed r r' f" and BIJ: "bij_betw f (Field r) (Field r')" +shows "embed r' r (inv_into (Field r) f)" +proof- + have "Field r' \ f ` (Field r)" + using BIJ by (auto simp add: bij_betw_def) + thus ?thesis using assms + by(auto simp add: inv_into_Field_embed) +qed + + + + + +subsection {* Given any two well-orders, one can be embedded in the other *} + + +text{* Here is an overview of the proof of of this fact, stated in theorem +@{text "wellorders_totally_ordered"}: + + Fix the well-orders @{text "r::'a rel"} and @{text "r'::'a' rel"}. + Attempt to define an embedding @{text "f::'a \ 'a'"} from @{text "r"} to @{text "r'"} in the + natural way by well-order recursion ("hoping" that @{text "Field r"} turns out to be smaller + than @{text "Field r'"}), but also record, at the recursive step, in a function + @{text "g::'a \ bool"}, the extra information of whether @{text "Field r'"} + gets exhausted or not. + + If @{text "Field r'"} does not get exhausted, then @{text "Field r"} is indeed smaller + and @{text "f"} is the desired embedding from @{text "r"} to @{text "r'"} + (lemma @{text "wellorders_totally_ordered_aux"}). + + Otherwise, it means that @{text "Field r'"} is the smaller one, and the inverse of + (the "good" segment of) @{text "f"} is the desired embedding from @{text "r'"} to @{text "r"} + (lemma @{text "wellorders_totally_ordered_aux2"}). +*} + + +lemma wellorders_totally_ordered_aux: +fixes r ::"'a rel" and r'::"'a' rel" and + f :: "'a \ 'a'" and a::'a +assumes WELL: "Well_order r" and WELL': "Well_order r'" and IN: "a \ Field r" and + IH: "\b \ rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))" and + NOT: "f ` (rel.underS r a) \ Field r'" and SUC: "f a = wo_rel.suc r' (f`(rel.underS r a))" +shows "bij_betw f (rel.under r a) (rel.under r' (f a))" +proof- + (* Preliminary facts *) + have Well: "wo_rel r" using WELL unfolding wo_rel_def . + hence Refl: "Refl r" using wo_rel.REFL[of r] by auto + have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto + have Well': "wo_rel r'" using WELL' unfolding wo_rel_def . + have OF: "wo_rel.ofilter r (rel.underS r a)" + by (auto simp add: Well wo_rel.underS_ofilter) + hence UN: "rel.underS r a = (\ b \ rel.underS r a. rel.under r b)" + using Well wo_rel.ofilter_under_UNION[of r "rel.underS r a"] by blast + (* Gather facts about elements of rel.underS r a *) + {fix b assume *: "b \ rel.underS r a" + hence t0: "(b,a) \ r \ b \ a" unfolding rel.underS_def by auto + have t1: "b \ Field r" + using * rel.underS_Field[of r a] by auto + have t2: "f`(rel.under r b) = rel.under r' (f b)" + using IH * by (auto simp add: bij_betw_def) + hence t3: "wo_rel.ofilter r' (f`(rel.under r b))" + using Well' by (auto simp add: wo_rel.under_ofilter) + have "f`(rel.under r b) \ Field r'" + using t2 by (auto simp add: rel.under_Field) + moreover + have "b \ rel.under r b" + using t1 by(auto simp add: Refl rel.Refl_under_in) + ultimately + have t4: "f b \ Field r'" by auto + have "f`(rel.under r b) = rel.under r' (f b) \ + wo_rel.ofilter r' (f`(rel.under r b)) \ + f b \ Field r'" + using t2 t3 t4 by auto + } + hence bFact: + "\b \ rel.underS r a. f`(rel.under r b) = rel.under r' (f b) \ + wo_rel.ofilter r' (f`(rel.under r b)) \ + f b \ Field r'" by blast + (* *) + have subField: "f`(rel.underS r a) \ Field r'" + using bFact by blast + (* *) + have OF': "wo_rel.ofilter r' (f`(rel.underS r a))" + proof- + have "f`(rel.underS r a) = f`(\ b \ rel.underS r a. rel.under r b)" + using UN by auto + also have "\ = (\ b \ rel.underS r a. f`(rel.under r b))" by blast + also have "\ = (\ b \ rel.underS r a. (rel.under r' (f b)))" + using bFact by auto + finally + have "f`(rel.underS r a) = (\ b \ rel.underS r a. (rel.under r' (f b)))" . + thus ?thesis + using Well' bFact + wo_rel.ofilter_UNION[of r' "rel.underS r a" "\ b. rel.under r' (f b)"] by fastforce + qed + (* *) + have "f`(rel.underS r a) \ rel.AboveS r' (f`(rel.underS r a)) = Field r'" + using Well' OF' by (auto simp add: wo_rel.ofilter_AboveS_Field) + hence NE: "rel.AboveS r' (f`(rel.underS r a)) \ {}" + using subField NOT by blast + (* Main proof *) + have INCL1: "f`(rel.underS r a) \ rel.underS r' (f a) " + proof(auto) + fix b assume *: "b \ rel.underS r a" + have "f b \ f a \ (f b, f a) \ r'" + using subField Well' SUC NE * + wo_rel.suc_greater[of r' "f`(rel.underS r a)" "f b"] by auto + thus "f b \ rel.underS r' (f a)" + unfolding rel.underS_def by simp + qed + (* *) + have INCL2: "rel.underS r' (f a) \ f`(rel.underS r a)" + proof + fix b' assume "b' \ rel.underS r' (f a)" + hence "b' \ f a \ (b', f a) \ r'" + unfolding rel.underS_def by simp + thus "b' \ f`(rel.underS r a)" + using Well' SUC NE OF' + wo_rel.suc_ofilter_in[of r' "f ` rel.underS r a" b'] by auto + qed + (* *) + have INJ: "inj_on f (rel.underS r a)" + proof- + have "\b \ rel.underS r a. inj_on f (rel.under r b)" + using IH by (auto simp add: bij_betw_def) + moreover + have "\b. wo_rel.ofilter r (rel.under r b)" + using Well by (auto simp add: wo_rel.under_ofilter) + ultimately show ?thesis + using WELL bFact UN + UNION_inj_on_ofilter[of r "rel.underS r a" "\b. rel.under r b" f] + by auto + qed + (* *) + have BIJ: "bij_betw f (rel.underS r a) (rel.underS r' (f a))" + unfolding bij_betw_def + using INJ INCL1 INCL2 by auto + (* *) + have "f a \ Field r'" + using Well' subField NE SUC + by (auto simp add: wo_rel.suc_inField) + thus ?thesis + using WELL WELL' IN BIJ under_underS_bij_betw[of r r' a f] by auto +qed + + +lemma wellorders_totally_ordered_aux2: +fixes r ::"'a rel" and r'::"'a' rel" and + f :: "'a \ 'a'" and g :: "'a \ bool" and a::'a +assumes WELL: "Well_order r" and WELL': "Well_order r'" and +MAIN1: + "\ a. (False \ g`(rel.underS r a) \ f`(rel.underS r a) \ Field r' + \ f a = wo_rel.suc r' (f`(rel.underS r a)) \ g a = True) + \ + (\(False \ (g`(rel.underS r a)) \ f`(rel.underS r a) \ Field r') + \ g a = False)" and +MAIN2: "\ a. a \ Field r \ False \ g`(rel.under r a) \ + bij_betw f (rel.under r a) (rel.under r' (f a))" and +Case: "a \ Field r \ False \ g`(rel.under r a)" +shows "\f'. embed r' r f'" +proof- + have Well: "wo_rel r" using WELL unfolding wo_rel_def . + hence Refl: "Refl r" using wo_rel.REFL[of r] by auto + have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto + have Antisym: "antisym r" using Well wo_rel.ANTISYM[of r] by auto + have Well': "wo_rel r'" using WELL' unfolding wo_rel_def . + (* *) + have 0: "rel.under r a = rel.underS r a \ {a}" + using Refl Case by(auto simp add: rel.Refl_under_underS) + (* *) + have 1: "g a = False" + proof- + {assume "g a \ False" + with 0 Case have "False \ g`(rel.underS r a)" by blast + with MAIN1 have "g a = False" by blast} + thus ?thesis by blast + qed + let ?A = "{a \ Field r. g a = False}" + let ?a = "(wo_rel.minim r ?A)" + (* *) + have 2: "?A \ {} \ ?A \ Field r" using Case 1 by blast + (* *) + have 3: "False \ g`(rel.underS r ?a)" + proof + assume "False \ g`(rel.underS r ?a)" + then obtain b where "b \ rel.underS r ?a" and 31: "g b = False" by auto + hence 32: "(b,?a) \ r \ b \ ?a" + by (auto simp add: rel.underS_def) + hence "b \ Field r" unfolding Field_def by auto + with 31 have "b \ ?A" by auto + hence "(?a,b) \ r" using wo_rel.minim_least 2 Well by fastforce + (* again: why worked without type annotations? *) + with 32 Antisym show False + by (auto simp add: antisym_def) + qed + have temp: "?a \ ?A" + using Well 2 wo_rel.minim_in[of r ?A] by auto + hence 4: "?a \ Field r" by auto + (* *) + have 5: "g ?a = False" using temp by blast + (* *) + have 6: "f`(rel.underS r ?a) = Field r'" + using MAIN1[of ?a] 3 5 by blast + (* *) + have 7: "\b \ rel.underS r ?a. bij_betw f (rel.under r b) (rel.under r' (f b))" + proof + fix b assume as: "b \ rel.underS r ?a" + moreover + have "wo_rel.ofilter r (rel.underS r ?a)" + using Well by (auto simp add: wo_rel.underS_ofilter) + ultimately + have "False \ g`(rel.under r b)" using 3 Well by (subst (asm) wo_rel.ofilter_def) fast+ + moreover have "b \ Field r" + unfolding Field_def using as by (auto simp add: rel.underS_def) + ultimately + show "bij_betw f (rel.under r b) (rel.under r' (f b))" + using MAIN2 by auto + qed + (* *) + have "embed r' r (inv_into (rel.underS r ?a) f)" + using WELL WELL' 7 4 6 inv_into_underS_embed[of r ?a f r'] by auto + thus ?thesis + unfolding embed_def by blast +qed + + +theorem wellorders_totally_ordered: +fixes r ::"'a rel" and r'::"'a' rel" +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "(\f. embed r r' f) \ (\f'. embed r' r f')" +proof- + (* Preliminary facts *) + have Well: "wo_rel r" using WELL unfolding wo_rel_def . + hence Refl: "Refl r" using wo_rel.REFL[of r] by auto + have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto + have Well': "wo_rel r'" using WELL' unfolding wo_rel_def . + (* Main proof *) + obtain H where H_def: "H = + (\h a. if False \ (snd o h)`(rel.underS r a) \ (fst o h)`(rel.underS r a) \ Field r' + then (wo_rel.suc r' ((fst o h)`(rel.underS r a)), True) + else (undefined, False))" by blast + have Adm: "wo_rel.adm_wo r H" + using Well + proof(unfold wo_rel.adm_wo_def, clarify) + fix h1::"'a \ 'a' * bool" and h2::"'a \ 'a' * bool" and x + assume "\y\rel.underS r x. h1 y = h2 y" + hence "\y\rel.underS r x. (fst o h1) y = (fst o h2) y \ + (snd o h1) y = (snd o h2) y" by auto + hence "(fst o h1)`(rel.underS r x) = (fst o h2)`(rel.underS r x) \ + (snd o h1)`(rel.underS r x) = (snd o h2)`(rel.underS r x)" + by (auto simp add: image_def) + thus "H h1 x = H h2 x" by (simp add: H_def del: not_False_in_image_Ball) + qed + (* More constant definitions: *) + obtain h::"'a \ 'a' * bool" and f::"'a \ 'a'" and g::"'a \ bool" + where h_def: "h = wo_rel.worec r H" and + f_def: "f = fst o h" and g_def: "g = snd o h" by blast + obtain test where test_def: + "test = (\ a. False \ (g`(rel.underS r a)) \ f`(rel.underS r a) \ Field r')" by blast + (* *) + have *: "\ a. h a = H h a" + using Adm Well wo_rel.worec_fixpoint[of r H] by (simp add: h_def) + have Main1: + "\ a. (test a \ f a = wo_rel.suc r' (f`(rel.underS r a)) \ g a = True) \ + (\(test a) \ g a = False)" + proof- (* How can I prove this withou fixing a? *) + fix a show "(test a \ f a = wo_rel.suc r' (f`(rel.underS r a)) \ g a = True) \ + (\(test a) \ g a = False)" + using *[of a] test_def f_def g_def H_def by auto + qed + (* *) + let ?phi = "\ a. a \ Field r \ False \ g`(rel.under r a) \ + bij_betw f (rel.under r a) (rel.under r' (f a))" + have Main2: "\ a. ?phi a" + proof- + fix a show "?phi a" + proof(rule wo_rel.well_order_induct[of r ?phi], + simp only: Well, clarify) + fix a + assume IH: "\b. b \ a \ (b,a) \ r \ ?phi b" and + *: "a \ Field r" and + **: "False \ g`(rel.under r a)" + have 1: "\b \ rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))" + proof(clarify) + fix b assume ***: "b \ rel.underS r a" + hence 0: "(b,a) \ r \ b \ a" unfolding rel.underS_def by auto + moreover have "b \ Field r" + using *** rel.underS_Field[of r a] by auto + moreover have "False \ g`(rel.under r b)" + using 0 ** Trans rel.under_incr[of r b a] by auto + ultimately show "bij_betw f (rel.under r b) (rel.under r' (f b))" + using IH by auto + qed + (* *) + have 21: "False \ g`(rel.underS r a)" + using ** rel.underS_subset_under[of r a] by auto + have 22: "g`(rel.under r a) \ {True}" using ** by auto + moreover have 23: "a \ rel.under r a" + using Refl * by (auto simp add: rel.Refl_under_in) + ultimately have 24: "g a = True" by blast + have 2: "f`(rel.underS r a) \ Field r'" + proof + assume "f`(rel.underS r a) = Field r'" + hence "g a = False" using Main1 test_def by blast + with 24 show False using ** by blast + qed + (* *) + have 3: "f a = wo_rel.suc r' (f`(rel.underS r a))" + using 21 2 Main1 test_def by blast + (* *) + show "bij_betw f (rel.under r a) (rel.under r' (f a))" + using WELL WELL' 1 2 3 * + wellorders_totally_ordered_aux[of r r' a f] by auto + qed + qed + (* *) + let ?chi = "(\ a. a \ Field r \ False \ g`(rel.under r a))" + show ?thesis + proof(cases "\a. ?chi a") + assume "\ (\a. ?chi a)" + hence "\a \ Field r. bij_betw f (rel.under r a) (rel.under r' (f a))" + using Main2 by blast + thus ?thesis unfolding embed_def by blast + next + assume "\a. ?chi a" + then obtain a where "?chi a" by blast + hence "\f'. embed r' r f'" + using wellorders_totally_ordered_aux2[of r r' g f a] + WELL WELL' Main1 Main2 test_def by blast + thus ?thesis by blast + qed +qed + + +subsection {* Uniqueness of embeddings *} + + +text{* Here we show a fact complementary to the one from the previous subsection -- namely, +that between any two well-orders there is {\em at most} one embedding, and is the one +definable by the expected well-order recursive equation. As a consequence, any two +embeddings of opposite directions are mutually inverse. *} + + +lemma embed_determined: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r r' f" and IN: "a \ Field r" +shows "f a = wo_rel.suc r' (f`(rel.underS r a))" +proof- + have "bij_betw f (rel.underS r a) (rel.underS r' (f a))" + using assms by (auto simp add: embed_underS) + hence "f`(rel.underS r a) = rel.underS r' (f a)" + by (auto simp add: bij_betw_def) + moreover + {have "f a \ Field r'" using IN + using EMB WELL embed_Field[of r r' f] by auto + hence "f a = wo_rel.suc r' (rel.underS r' (f a))" + using WELL' by (auto simp add: wo_rel_def wo_rel.suc_underS) + } + ultimately show ?thesis by simp +qed + + +lemma embed_unique: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMBf: "embed r r' f" and EMBg: "embed r r' g" +shows "a \ Field r \ f a = g a" +proof(rule wo_rel.well_order_induct[of r], auto simp add: WELL wo_rel_def) + fix a + assume IH: "\b. b \ a \ (b,a): r \ b \ Field r \ f b = g b" and + *: "a \ Field r" + hence "\b \ rel.underS r a. f b = g b" + unfolding rel.underS_def by (auto simp add: Field_def) + hence "f`(rel.underS r a) = g`(rel.underS r a)" by force + thus "f a = g a" + using assms * embed_determined[of r r' f a] embed_determined[of r r' g a] by auto +qed + + +lemma embed_bothWays_inverse: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r r' f" and EMB': "embed r' r f'" +shows "(\a \ Field r. f'(f a) = a) \ (\a' \ Field r'. f(f' a') = a')" +proof- + have "embed r r (f' o f)" using assms + by(auto simp add: comp_embed) + moreover have "embed r r id" using assms + by (auto simp add: id_embed) + ultimately have "\a \ Field r. f'(f a) = a" + using assms embed_unique[of r r "f' o f" id] id_def by auto + moreover + {have "embed r' r' (f o f')" using assms + by(auto simp add: comp_embed) + moreover have "embed r' r' id" using assms + by (auto simp add: id_embed) + ultimately have "\a' \ Field r'. f(f' a') = a'" + using assms embed_unique[of r' r' "f o f'" id] id_def by auto + } + ultimately show ?thesis by blast +qed + + +lemma embed_bothWays_bij_betw: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r r' f" and EMB': "embed r' r g" +shows "bij_betw f (Field r) (Field r')" +proof- + let ?A = "Field r" let ?A' = "Field r'" + have "embed r r (g o f) \ embed r' r' (f o g)" + using assms by (auto simp add: comp_embed) + hence 1: "(\a \ ?A. g(f a) = a) \ (\a' \ ?A'. f(g a') = a')" + using WELL id_embed[of r] embed_unique[of r r "g o f" id] + WELL' id_embed[of r'] embed_unique[of r' r' "f o g" id] + id_def by auto + have 2: "(\a \ ?A. f a \ ?A') \ (\a' \ ?A'. g a' \ ?A)" + using assms embed_Field[of r r' f] embed_Field[of r' r g] by blast + (* *) + show ?thesis + proof(unfold bij_betw_def inj_on_def, auto simp add: 2) + fix a b assume *: "a \ ?A" "b \ ?A" and **: "f a = f b" + have "a = g(f a) \ b = g(f b)" using * 1 by auto + with ** show "a = b" by auto + next + fix a' assume *: "a' \ ?A'" + hence "g a' \ ?A \ f(g a') = a'" using 1 2 by auto + thus "a' \ f ` ?A" by force + qed +qed + + +lemma embed_bothWays_iso: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r r' f" and EMB': "embed r' r g" +shows "iso r r' f" +unfolding iso_def using assms by (auto simp add: embed_bothWays_bij_betw) + + +subsection {* More properties of embeddings, strict embeddings and isomorphisms *} + + +lemma embed_bothWays_Field_bij_betw: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and + EMB: "embed r r' f" and EMB': "embed r' r f'" +shows "bij_betw f (Field r) (Field r')" +proof- + have "(\a \ Field r. f'(f a) = a) \ (\a' \ Field r'. f(f' a') = a')" + using assms by (auto simp add: embed_bothWays_inverse) + moreover + have "f`(Field r) \ Field r' \ f' ` (Field r') \ Field r" + using assms by (auto simp add: embed_Field) + ultimately + show ?thesis using bij_betw_byWitness[of "Field r" f' f "Field r'"] by auto +qed + + +lemma embedS_comp_embed: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" + and EMB: "embedS r r' f" and EMB': "embed r' r'' f'" +shows "embedS r r'' (f' o f)" +proof- + let ?g = "(f' o f)" let ?h = "inv_into (Field r) ?g" + have 1: "embed r r' f \ \ (bij_betw f (Field r) (Field r'))" + using EMB by (auto simp add: embedS_def) + hence 2: "embed r r'' ?g" + using WELL EMB' comp_embed[of r r' f r'' f'] by auto + moreover + {assume "bij_betw ?g (Field r) (Field r'')" + hence "embed r'' r ?h" using 2 WELL + by (auto simp add: inv_into_Field_embed_bij_betw) + hence "embed r' r (?h o f')" using WELL' EMB' + by (auto simp add: comp_embed) + hence "bij_betw f (Field r) (Field r')" using WELL WELL' 1 + by (auto simp add: embed_bothWays_Field_bij_betw) + with 1 have False by blast + } + ultimately show ?thesis unfolding embedS_def by auto +qed + + +lemma embed_comp_embedS: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" + and EMB: "embed r r' f" and EMB': "embedS r' r'' f'" +shows "embedS r r'' (f' o f)" +proof- + let ?g = "(f' o f)" let ?h = "inv_into (Field r) ?g" + have 1: "embed r' r'' f' \ \ (bij_betw f' (Field r') (Field r''))" + using EMB' by (auto simp add: embedS_def) + hence 2: "embed r r'' ?g" + using WELL EMB comp_embed[of r r' f r'' f'] by auto + moreover + {assume "bij_betw ?g (Field r) (Field r'')" + hence "embed r'' r ?h" using 2 WELL + by (auto simp add: inv_into_Field_embed_bij_betw) + hence "embed r'' r' (f o ?h)" using WELL'' EMB + by (auto simp add: comp_embed) + hence "bij_betw f' (Field r') (Field r'')" using WELL' WELL'' 1 + by (auto simp add: embed_bothWays_Field_bij_betw) + with 1 have False by blast + } + ultimately show ?thesis unfolding embedS_def by auto +qed + + +lemma embed_comp_iso: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" + and EMB: "embed r r' f" and EMB': "iso r' r'' f'" +shows "embed r r'' (f' o f)" +using assms unfolding iso_def +by (auto simp add: comp_embed) + + +lemma iso_comp_embed: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" + and EMB: "iso r r' f" and EMB': "embed r' r'' f'" +shows "embed r r'' (f' o f)" +using assms unfolding iso_def +by (auto simp add: comp_embed) + + +lemma embedS_comp_iso: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" + and EMB: "embedS r r' f" and EMB': "iso r' r'' f'" +shows "embedS r r'' (f' o f)" +using assms unfolding iso_def +by (auto simp add: embedS_comp_embed) + + +lemma iso_comp_embedS: +assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" + and EMB: "iso r r' f" and EMB': "embedS r' r'' f'" +shows "embedS r r'' (f' o f)" +using assms unfolding iso_def using embed_comp_embedS +by (auto simp add: embed_comp_embedS) + + +lemma embedS_Field: +assumes WELL: "Well_order r" and EMB: "embedS r r' f" +shows "f ` (Field r) < Field r'" +proof- + have "f`(Field r) \ Field r'" using assms + by (auto simp add: embed_Field embedS_def) + moreover + {have "inj_on f (Field r)" using assms + by (auto simp add: embedS_def embed_inj_on) + hence "f`(Field r) \ Field r'" using EMB + by (auto simp add: embedS_def bij_betw_def) + } + ultimately show ?thesis by blast +qed + + +lemma embedS_iff: +assumes WELL: "Well_order r" and ISO: "embed r r' f" +shows "embedS r r' f = (f ` (Field r) < Field r')" +proof + assume "embedS r r' f" + thus "f ` Field r \ Field r'" + using WELL by (auto simp add: embedS_Field) +next + assume "f ` Field r \ Field r'" + hence "\ bij_betw f (Field r) (Field r')" + unfolding bij_betw_def by blast + thus "embedS r r' f" unfolding embedS_def + using ISO by auto +qed + + +lemma iso_Field: +"iso r r' f \ f ` (Field r) = Field r'" +using assms by (auto simp add: iso_def bij_betw_def) + + +lemma iso_iff: +assumes "Well_order r" +shows "iso r r' f = (embed r r' f \ f ` (Field r) = Field r')" +proof + assume "iso r r' f" + thus "embed r r' f \ f ` (Field r) = Field r'" + by (auto simp add: iso_Field iso_def) +next + assume *: "embed r r' f \ f ` Field r = Field r'" + hence "inj_on f (Field r)" using assms by (auto simp add: embed_inj_on) + with * have "bij_betw f (Field r) (Field r')" + unfolding bij_betw_def by simp + with * show "iso r r' f" unfolding iso_def by auto +qed + + +lemma iso_iff2: +assumes "Well_order r" +shows "iso r r' f = (bij_betw f (Field r) (Field r') \ + (\a \ Field r. \b \ Field r. + (((a,b) \ r) = ((f a, f b) \ r'))))" +using assms +proof(auto simp add: iso_def) + fix a b + assume "embed r r' f" + hence "compat r r' f" using embed_compat[of r] by auto + moreover assume "(a,b) \ r" + ultimately show "(f a, f b) \ r'" using compat_def[of r] by auto +next + let ?f' = "inv_into (Field r) f" + assume "embed r r' f" and 1: "bij_betw f (Field r) (Field r')" + hence "embed r' r ?f'" using assms + by (auto simp add: inv_into_Field_embed_bij_betw) + hence 2: "compat r' r ?f'" using embed_compat[of r'] by auto + fix a b assume *: "a \ Field r" "b \ Field r" and **: "(f a,f b) \ r'" + hence "?f'(f a) = a \ ?f'(f b) = b" using 1 + by (auto simp add: bij_betw_inv_into_left) + thus "(a,b) \ r" using ** 2 compat_def[of r' r ?f'] by fastforce +next + assume *: "bij_betw f (Field r) (Field r')" and + **: "\a\Field r. \b\Field r. ((a, b) \ r) = ((f a, f b) \ r')" + have 1: "\ a. rel.under r a \ Field r \ rel.under r' (f a) \ Field r'" + by (auto simp add: rel.under_Field) + have 2: "inj_on f (Field r)" using * by (auto simp add: bij_betw_def) + {fix a assume ***: "a \ Field r" + have "bij_betw f (rel.under r a) (rel.under r' (f a))" + proof(unfold bij_betw_def, auto) + show "inj_on f (rel.under r a)" + using 1 2 by (auto simp add: subset_inj_on) + next + fix b assume "b \ rel.under r a" + hence "a \ Field r \ b \ Field r \ (b,a) \ r" + unfolding rel.under_def by (auto simp add: Field_def Range_def Domain_def) + with 1 ** show "f b \ rel.under r' (f a)" + unfolding rel.under_def by auto + next + fix b' assume "b' \ rel.under r' (f a)" + hence 3: "(b',f a) \ r'" unfolding rel.under_def by simp + hence "b' \ Field r'" unfolding Field_def by auto + with * obtain b where "b \ Field r \ f b = b'" + unfolding bij_betw_def by force + with 3 ** *** + show "b' \ f ` (rel.under r a)" unfolding rel.under_def by blast + qed + } + thus "embed r r' f" unfolding embed_def using * by auto +qed + + +lemma iso_iff3: +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "iso r r' f = (bij_betw f (Field r) (Field r') \ compat r r' f)" +proof + assume "iso r r' f" + thus "bij_betw f (Field r) (Field r') \ compat r r' f" + unfolding compat_def using WELL by (auto simp add: iso_iff2 Field_def) +next + have Well: "wo_rel r \ wo_rel r'" using WELL WELL' + by (auto simp add: wo_rel_def) + assume *: "bij_betw f (Field r) (Field r') \ compat r r' f" + thus "iso r r' f" + unfolding "compat_def" using assms + proof(auto simp add: iso_iff2) + fix a b assume **: "a \ Field r" "b \ Field r" and + ***: "(f a, f b) \ r'" + {assume "(b,a) \ r \ b = a" + hence "(b,a): r"using Well ** wo_rel.REFL[of r] refl_on_def[of _ r] by blast + hence "(f b, f a) \ r'" using * unfolding compat_def by auto + hence "f a = f b" + using Well *** wo_rel.ANTISYM[of r'] antisym_def[of r'] by blast + hence "a = b" using * ** unfolding bij_betw_def inj_on_def by auto + hence "(a,b) \ r" using Well ** wo_rel.REFL[of r] refl_on_def[of _ r] by blast + } + thus "(a,b) \ r" + using Well ** wo_rel.TOTAL[of r] total_on_def[of _ r] by blast + qed +qed + + + +end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellorder_Embedding_LFP.thy --- a/src/HOL/Cardinals/Wellorder_Embedding_LFP.thy Mon Nov 18 18:04:45 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1146 +0,0 @@ -(* Title: HOL/Cardinals/Wellorder_Embedding_LFP.thy - Author: Andrei Popescu, TU Muenchen - Copyright 2012 - -Well-order embeddings (LFP). -*) - -header {* Well-Order Embeddings (LFP) *} - -theory Wellorder_Embedding_LFP -imports "~~/src/HOL/Library/Zorn" Fun_More_LFP Wellorder_Relation_LFP -begin - - -text{* In this section, we introduce well-order {\em embeddings} and {\em isomorphisms} and -prove their basic properties. The notion of embedding is considered from the point -of view of the theory of ordinals, and therefore requires the source to be injected -as an {\em initial segment} (i.e., {\em order filter}) of the target. A main result -of this section is the existence of embeddings (in one direction or another) between -any two well-orders, having as a consequence the fact that, given any two sets on -any two types, one is smaller than (i.e., can be injected into) the other. *} - - -subsection {* Auxiliaries *} - -lemma UNION_inj_on_ofilter: -assumes WELL: "Well_order r" and - OF: "\ i. i \ I \ wo_rel.ofilter r (A i)" and - INJ: "\ i. i \ I \ inj_on f (A i)" -shows "inj_on f (\ i \ I. A i)" -proof- - have "wo_rel r" using WELL by (simp add: wo_rel_def) - hence "\ i j. \i \ I; j \ I\ \ A i <= A j \ A j <= A i" - using wo_rel.ofilter_linord[of r] OF by blast - with WELL INJ show ?thesis - by (auto simp add: inj_on_UNION_chain) -qed - - -lemma under_underS_bij_betw: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - IN: "a \ Field r" and IN': "f a \ Field r'" and - BIJ: "bij_betw f (rel.underS r a) (rel.underS r' (f a))" -shows "bij_betw f (rel.under r a) (rel.under r' (f a))" -proof- - have "a \ rel.underS r a \ f a \ rel.underS r' (f a)" - unfolding rel.underS_def by auto - moreover - {have "Refl r \ Refl r'" using WELL WELL' - by (auto simp add: order_on_defs) - hence "rel.under r a = rel.underS r a \ {a} \ - rel.under r' (f a) = rel.underS r' (f a) \ {f a}" - using IN IN' by(auto simp add: rel.Refl_under_underS) - } - ultimately show ?thesis - using BIJ notIn_Un_bij_betw[of a "rel.underS r a" f "rel.underS r' (f a)"] by auto -qed - - - -subsection {* (Well-order) embeddings, strict embeddings, isomorphisms and order-compatible -functions *} - - -text{* Standardly, a function is an embedding of a well-order in another if it injectively and -order-compatibly maps the former into an order filter of the latter. -Here we opt for a more succinct definition (operator @{text "embed"}), -asking that, for any element in the source, the function should be a bijection -between the set of strict lower bounds of that element -and the set of strict lower bounds of its image. (Later we prove equivalence with -the standard definition -- lemma @{text "embed_iff_compat_inj_on_ofilter"}.) -A {\em strict embedding} (operator @{text "embedS"}) is a non-bijective embedding -and an isomorphism (operator @{text "iso"}) is a bijective embedding. *} - - -definition embed :: "'a rel \ 'a' rel \ ('a \ 'a') \ bool" -where -"embed r r' f \ \a \ Field r. bij_betw f (rel.under r a) (rel.under r' (f a))" - - -lemmas embed_defs = embed_def embed_def[abs_def] - - -text {* Strict embeddings: *} - -definition embedS :: "'a rel \ 'a' rel \ ('a \ 'a') \ bool" -where -"embedS r r' f \ embed r r' f \ \ bij_betw f (Field r) (Field r')" - - -lemmas embedS_defs = embedS_def embedS_def[abs_def] - - -definition iso :: "'a rel \ 'a' rel \ ('a \ 'a') \ bool" -where -"iso r r' f \ embed r r' f \ bij_betw f (Field r) (Field r')" - - -lemmas iso_defs = iso_def iso_def[abs_def] - - -definition compat :: "'a rel \ 'a' rel \ ('a \ 'a') \ bool" -where -"compat r r' f \ \a b. (a,b) \ r \ (f a, f b) \ r'" - - -lemma compat_wf: -assumes CMP: "compat r r' f" and WF: "wf r'" -shows "wf r" -proof- - have "r \ inv_image r' f" - unfolding inv_image_def using CMP - by (auto simp add: compat_def) - with WF show ?thesis - using wf_inv_image[of r' f] wf_subset[of "inv_image r' f"] by auto -qed - - -lemma id_embed: "embed r r id" -by(auto simp add: id_def embed_def bij_betw_def) - - -lemma id_iso: "iso r r id" -by(auto simp add: id_def embed_def iso_def bij_betw_def) - - -lemma embed_in_Field: -assumes WELL: "Well_order r" and - EMB: "embed r r' f" and IN: "a \ Field r" -shows "f a \ Field r'" -proof- - have Well: "wo_rel r" - using WELL by (auto simp add: wo_rel_def) - hence 1: "Refl r" - by (auto simp add: wo_rel.REFL) - hence "a \ rel.under r a" using IN rel.Refl_under_in by fastforce - hence "f a \ rel.under r' (f a)" - using EMB IN by (auto simp add: embed_def bij_betw_def) - thus ?thesis unfolding Field_def - by (auto simp: rel.under_def) -qed - - -lemma comp_embed: -assumes WELL: "Well_order r" and - EMB: "embed r r' f" and EMB': "embed r' r'' f'" -shows "embed r r'' (f' o f)" -proof(unfold embed_def, auto) - fix a assume *: "a \ Field r" - hence "bij_betw f (rel.under r a) (rel.under r' (f a))" - using embed_def[of r] EMB by auto - moreover - {have "f a \ Field r'" - using EMB WELL * by (auto simp add: embed_in_Field) - hence "bij_betw f' (rel.under r' (f a)) (rel.under r'' (f' (f a)))" - using embed_def[of r'] EMB' by auto - } - ultimately - show "bij_betw (f' \ f) (rel.under r a) (rel.under r'' (f'(f a)))" - by(auto simp add: bij_betw_trans) -qed - - -lemma comp_iso: -assumes WELL: "Well_order r" and - EMB: "iso r r' f" and EMB': "iso r' r'' f'" -shows "iso r r'' (f' o f)" -using assms unfolding iso_def -by (auto simp add: comp_embed bij_betw_trans) - - -text{* That @{text "embedS"} is also preserved by function composition shall be proved only later. *} - - -lemma embed_Field: -"\Well_order r; embed r r' f\ \ f`(Field r) \ Field r'" -by (auto simp add: embed_in_Field) - - -lemma embed_preserves_ofilter: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r r' f" and OF: "wo_rel.ofilter r A" -shows "wo_rel.ofilter r' (f`A)" -proof- - (* Preliminary facts *) - from WELL have Well: "wo_rel r" unfolding wo_rel_def . - from WELL' have Well': "wo_rel r'" unfolding wo_rel_def . - from OF have 0: "A \ Field r" by(auto simp add: Well wo_rel.ofilter_def) - (* Main proof *) - show ?thesis using Well' WELL EMB 0 embed_Field[of r r' f] - proof(unfold wo_rel.ofilter_def, auto simp add: image_def) - fix a b' - assume *: "a \ A" and **: "b' \ rel.under r' (f a)" - hence "a \ Field r" using 0 by auto - hence "bij_betw f (rel.under r a) (rel.under r' (f a))" - using * EMB by (auto simp add: embed_def) - hence "f`(rel.under r a) = rel.under r' (f a)" - by (simp add: bij_betw_def) - with ** image_def[of f "rel.under r a"] obtain b where - 1: "b \ rel.under r a \ b' = f b" by blast - hence "b \ A" using Well * OF - by (auto simp add: wo_rel.ofilter_def) - with 1 show "\b \ A. b' = f b" by blast - qed -qed - - -lemma embed_Field_ofilter: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r r' f" -shows "wo_rel.ofilter r' (f`(Field r))" -proof- - have "wo_rel.ofilter r (Field r)" - using WELL by (auto simp add: wo_rel_def wo_rel.Field_ofilter) - with WELL WELL' EMB - show ?thesis by (auto simp add: embed_preserves_ofilter) -qed - - -lemma embed_compat: -assumes EMB: "embed r r' f" -shows "compat r r' f" -proof(unfold compat_def, clarify) - fix a b - assume *: "(a,b) \ r" - hence 1: "b \ Field r" using Field_def[of r] by blast - have "a \ rel.under r b" - using * rel.under_def[of r] by simp - hence "f a \ rel.under r' (f b)" - using EMB embed_def[of r r' f] - bij_betw_def[of f "rel.under r b" "rel.under r' (f b)"] - image_def[of f "rel.under r b"] 1 by auto - thus "(f a, f b) \ r'" - by (auto simp add: rel.under_def) -qed - - -lemma embed_inj_on: -assumes WELL: "Well_order r" and EMB: "embed r r' f" -shows "inj_on f (Field r)" -proof(unfold inj_on_def, clarify) - (* Preliminary facts *) - from WELL have Well: "wo_rel r" unfolding wo_rel_def . - with wo_rel.TOTAL[of r] - have Total: "Total r" by simp - from Well wo_rel.REFL[of r] - have Refl: "Refl r" by simp - (* Main proof *) - fix a b - assume *: "a \ Field r" and **: "b \ Field r" and - ***: "f a = f b" - hence 1: "a \ Field r \ b \ Field r" - unfolding Field_def by auto - {assume "(a,b) \ r" - hence "a \ rel.under r b \ b \ rel.under r b" - using Refl by(auto simp add: rel.under_def refl_on_def) - hence "a = b" - using EMB 1 *** - by (auto simp add: embed_def bij_betw_def inj_on_def) - } - moreover - {assume "(b,a) \ r" - hence "a \ rel.under r a \ b \ rel.under r a" - using Refl by(auto simp add: rel.under_def refl_on_def) - hence "a = b" - using EMB 1 *** - by (auto simp add: embed_def bij_betw_def inj_on_def) - } - ultimately - show "a = b" using Total 1 - by (auto simp add: total_on_def) -qed - - -lemma embed_underS: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r r' f" and IN: "a \ Field r" -shows "bij_betw f (rel.underS r a) (rel.underS r' (f a))" -proof- - have "bij_betw f (rel.under r a) (rel.under r' (f a))" - using assms by (auto simp add: embed_def) - moreover - {have "f a \ Field r'" using assms embed_Field[of r r' f] by auto - hence "rel.under r a = rel.underS r a \ {a} \ - rel.under r' (f a) = rel.underS r' (f a) \ {f a}" - using assms by (auto simp add: order_on_defs rel.Refl_under_underS) - } - moreover - {have "a \ rel.underS r a \ f a \ rel.underS r' (f a)" - unfolding rel.underS_def by blast - } - ultimately show ?thesis - by (auto simp add: notIn_Un_bij_betw3) -qed - - -lemma embed_iff_compat_inj_on_ofilter: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "embed r r' f = (compat r r' f \ inj_on f (Field r) \ wo_rel.ofilter r' (f`(Field r)))" -using assms -proof(auto simp add: embed_compat embed_inj_on embed_Field_ofilter, - unfold embed_def, auto) (* get rid of one implication *) - fix a - assume *: "inj_on f (Field r)" and - **: "compat r r' f" and - ***: "wo_rel.ofilter r' (f`(Field r))" and - ****: "a \ Field r" - (* Preliminary facts *) - have Well: "wo_rel r" - using WELL wo_rel_def[of r] by simp - hence Refl: "Refl r" - using wo_rel.REFL[of r] by simp - have Total: "Total r" - using Well wo_rel.TOTAL[of r] by simp - have Well': "wo_rel r'" - using WELL' wo_rel_def[of r'] by simp - hence Antisym': "antisym r'" - using wo_rel.ANTISYM[of r'] by simp - have "(a,a) \ r" - using **** Well wo_rel.REFL[of r] - refl_on_def[of _ r] by auto - hence "(f a, f a) \ r'" - using ** by(auto simp add: compat_def) - hence 0: "f a \ Field r'" - unfolding Field_def by auto - have "f a \ f`(Field r)" - using **** by auto - hence 2: "rel.under r' (f a) \ f`(Field r)" - using Well' *** wo_rel.ofilter_def[of r' "f`(Field r)"] by fastforce - (* Main proof *) - show "bij_betw f (rel.under r a) (rel.under r' (f a))" - proof(unfold bij_betw_def, auto) - show "inj_on f (rel.under r a)" - using * - by (auto simp add: rel.under_Field subset_inj_on) - next - fix b assume "b \ rel.under r a" - thus "f b \ rel.under r' (f a)" - unfolding rel.under_def using ** - by (auto simp add: compat_def) - next - fix b' assume *****: "b' \ rel.under r' (f a)" - hence "b' \ f`(Field r)" - using 2 by auto - with Field_def[of r] obtain b where - 3: "b \ Field r" and 4: "b' = f b" by auto - have "(b,a): r" - proof- - {assume "(a,b) \ r" - with ** 4 have "(f a, b'): r'" - by (auto simp add: compat_def) - with ***** Antisym' have "f a = b'" - by(auto simp add: rel.under_def antisym_def) - with 3 **** 4 * have "a = b" - by(auto simp add: inj_on_def) - } - moreover - {assume "a = b" - hence "(b,a) \ r" using Refl **** 3 - by (auto simp add: refl_on_def) - } - ultimately - show ?thesis using Total **** 3 by (fastforce simp add: total_on_def) - qed - with 4 show "b' \ f`(rel.under r a)" - unfolding rel.under_def by auto - qed -qed - - -lemma inv_into_ofilter_embed: -assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and - BIJ: "\b \ A. bij_betw f (rel.under r b) (rel.under r' (f b))" and - IMAGE: "f ` A = Field r'" -shows "embed r' r (inv_into A f)" -proof- - (* Preliminary facts *) - have Well: "wo_rel r" - using WELL wo_rel_def[of r] by simp - have Refl: "Refl r" - using Well wo_rel.REFL[of r] by simp - have Total: "Total r" - using Well wo_rel.TOTAL[of r] by simp - (* Main proof *) - have 1: "bij_betw f A (Field r')" - proof(unfold bij_betw_def inj_on_def, auto simp add: IMAGE) - fix b1 b2 - assume *: "b1 \ A" and **: "b2 \ A" and - ***: "f b1 = f b2" - have 11: "b1 \ Field r \ b2 \ Field r" - using * ** Well OF by (auto simp add: wo_rel.ofilter_def) - moreover - {assume "(b1,b2) \ r" - hence "b1 \ rel.under r b2 \ b2 \ rel.under r b2" - unfolding rel.under_def using 11 Refl - by (auto simp add: refl_on_def) - hence "b1 = b2" using BIJ * ** *** - by (auto simp add: bij_betw_def inj_on_def) - } - moreover - {assume "(b2,b1) \ r" - hence "b1 \ rel.under r b1 \ b2 \ rel.under r b1" - unfolding rel.under_def using 11 Refl - by (auto simp add: refl_on_def) - hence "b1 = b2" using BIJ * ** *** - by (auto simp add: bij_betw_def inj_on_def) - } - ultimately - show "b1 = b2" - using Total by (auto simp add: total_on_def) - qed - (* *) - let ?f' = "(inv_into A f)" - (* *) - have 2: "\b \ A. bij_betw ?f' (rel.under r' (f b)) (rel.under r b)" - proof(clarify) - fix b assume *: "b \ A" - hence "rel.under r b \ A" - using Well OF by(auto simp add: wo_rel.ofilter_def) - moreover - have "f ` (rel.under r b) = rel.under r' (f b)" - using * BIJ by (auto simp add: bij_betw_def) - ultimately - show "bij_betw ?f' (rel.under r' (f b)) (rel.under r b)" - using 1 by (auto simp add: bij_betw_inv_into_subset) - qed - (* *) - have 3: "\b' \ Field r'. bij_betw ?f' (rel.under r' b') (rel.under r (?f' b'))" - proof(clarify) - fix b' assume *: "b' \ Field r'" - have "b' = f (?f' b')" using * 1 - by (auto simp add: bij_betw_inv_into_right) - moreover - {obtain b where 31: "b \ A" and "f b = b'" using IMAGE * by force - hence "?f' b' = b" using 1 by (auto simp add: bij_betw_inv_into_left) - with 31 have "?f' b' \ A" by auto - } - ultimately - show "bij_betw ?f' (rel.under r' b') (rel.under r (?f' b'))" - using 2 by auto - qed - (* *) - thus ?thesis unfolding embed_def . -qed - - -lemma inv_into_underS_embed: -assumes WELL: "Well_order r" and - BIJ: "\b \ rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))" and - IN: "a \ Field r" and - IMAGE: "f ` (rel.underS r a) = Field r'" -shows "embed r' r (inv_into (rel.underS r a) f)" -using assms -by(auto simp add: wo_rel_def wo_rel.underS_ofilter inv_into_ofilter_embed) - - -lemma inv_into_Field_embed: -assumes WELL: "Well_order r" and EMB: "embed r r' f" and - IMAGE: "Field r' \ f ` (Field r)" -shows "embed r' r (inv_into (Field r) f)" -proof- - have "(\b \ Field r. bij_betw f (rel.under r b) (rel.under r' (f b)))" - using EMB by (auto simp add: embed_def) - moreover - have "f ` (Field r) \ Field r'" - using EMB WELL by (auto simp add: embed_Field) - ultimately - show ?thesis using assms - by(auto simp add: wo_rel_def wo_rel.Field_ofilter inv_into_ofilter_embed) -qed - - -lemma inv_into_Field_embed_bij_betw: -assumes WELL: "Well_order r" and - EMB: "embed r r' f" and BIJ: "bij_betw f (Field r) (Field r')" -shows "embed r' r (inv_into (Field r) f)" -proof- - have "Field r' \ f ` (Field r)" - using BIJ by (auto simp add: bij_betw_def) - thus ?thesis using assms - by(auto simp add: inv_into_Field_embed) -qed - - - - - -subsection {* Given any two well-orders, one can be embedded in the other *} - - -text{* Here is an overview of the proof of of this fact, stated in theorem -@{text "wellorders_totally_ordered"}: - - Fix the well-orders @{text "r::'a rel"} and @{text "r'::'a' rel"}. - Attempt to define an embedding @{text "f::'a \ 'a'"} from @{text "r"} to @{text "r'"} in the - natural way by well-order recursion ("hoping" that @{text "Field r"} turns out to be smaller - than @{text "Field r'"}), but also record, at the recursive step, in a function - @{text "g::'a \ bool"}, the extra information of whether @{text "Field r'"} - gets exhausted or not. - - If @{text "Field r'"} does not get exhausted, then @{text "Field r"} is indeed smaller - and @{text "f"} is the desired embedding from @{text "r"} to @{text "r'"} - (lemma @{text "wellorders_totally_ordered_aux"}). - - Otherwise, it means that @{text "Field r'"} is the smaller one, and the inverse of - (the "good" segment of) @{text "f"} is the desired embedding from @{text "r'"} to @{text "r"} - (lemma @{text "wellorders_totally_ordered_aux2"}). -*} - - -lemma wellorders_totally_ordered_aux: -fixes r ::"'a rel" and r'::"'a' rel" and - f :: "'a \ 'a'" and a::'a -assumes WELL: "Well_order r" and WELL': "Well_order r'" and IN: "a \ Field r" and - IH: "\b \ rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))" and - NOT: "f ` (rel.underS r a) \ Field r'" and SUC: "f a = wo_rel.suc r' (f`(rel.underS r a))" -shows "bij_betw f (rel.under r a) (rel.under r' (f a))" -proof- - (* Preliminary facts *) - have Well: "wo_rel r" using WELL unfolding wo_rel_def . - hence Refl: "Refl r" using wo_rel.REFL[of r] by auto - have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto - have Well': "wo_rel r'" using WELL' unfolding wo_rel_def . - have OF: "wo_rel.ofilter r (rel.underS r a)" - by (auto simp add: Well wo_rel.underS_ofilter) - hence UN: "rel.underS r a = (\ b \ rel.underS r a. rel.under r b)" - using Well wo_rel.ofilter_under_UNION[of r "rel.underS r a"] by blast - (* Gather facts about elements of rel.underS r a *) - {fix b assume *: "b \ rel.underS r a" - hence t0: "(b,a) \ r \ b \ a" unfolding rel.underS_def by auto - have t1: "b \ Field r" - using * rel.underS_Field[of r a] by auto - have t2: "f`(rel.under r b) = rel.under r' (f b)" - using IH * by (auto simp add: bij_betw_def) - hence t3: "wo_rel.ofilter r' (f`(rel.under r b))" - using Well' by (auto simp add: wo_rel.under_ofilter) - have "f`(rel.under r b) \ Field r'" - using t2 by (auto simp add: rel.under_Field) - moreover - have "b \ rel.under r b" - using t1 by(auto simp add: Refl rel.Refl_under_in) - ultimately - have t4: "f b \ Field r'" by auto - have "f`(rel.under r b) = rel.under r' (f b) \ - wo_rel.ofilter r' (f`(rel.under r b)) \ - f b \ Field r'" - using t2 t3 t4 by auto - } - hence bFact: - "\b \ rel.underS r a. f`(rel.under r b) = rel.under r' (f b) \ - wo_rel.ofilter r' (f`(rel.under r b)) \ - f b \ Field r'" by blast - (* *) - have subField: "f`(rel.underS r a) \ Field r'" - using bFact by blast - (* *) - have OF': "wo_rel.ofilter r' (f`(rel.underS r a))" - proof- - have "f`(rel.underS r a) = f`(\ b \ rel.underS r a. rel.under r b)" - using UN by auto - also have "\ = (\ b \ rel.underS r a. f`(rel.under r b))" by blast - also have "\ = (\ b \ rel.underS r a. (rel.under r' (f b)))" - using bFact by auto - finally - have "f`(rel.underS r a) = (\ b \ rel.underS r a. (rel.under r' (f b)))" . - thus ?thesis - using Well' bFact - wo_rel.ofilter_UNION[of r' "rel.underS r a" "\ b. rel.under r' (f b)"] by fastforce - qed - (* *) - have "f`(rel.underS r a) \ rel.AboveS r' (f`(rel.underS r a)) = Field r'" - using Well' OF' by (auto simp add: wo_rel.ofilter_AboveS_Field) - hence NE: "rel.AboveS r' (f`(rel.underS r a)) \ {}" - using subField NOT by blast - (* Main proof *) - have INCL1: "f`(rel.underS r a) \ rel.underS r' (f a) " - proof(auto) - fix b assume *: "b \ rel.underS r a" - have "f b \ f a \ (f b, f a) \ r'" - using subField Well' SUC NE * - wo_rel.suc_greater[of r' "f`(rel.underS r a)" "f b"] by auto - thus "f b \ rel.underS r' (f a)" - unfolding rel.underS_def by simp - qed - (* *) - have INCL2: "rel.underS r' (f a) \ f`(rel.underS r a)" - proof - fix b' assume "b' \ rel.underS r' (f a)" - hence "b' \ f a \ (b', f a) \ r'" - unfolding rel.underS_def by simp - thus "b' \ f`(rel.underS r a)" - using Well' SUC NE OF' - wo_rel.suc_ofilter_in[of r' "f ` rel.underS r a" b'] by auto - qed - (* *) - have INJ: "inj_on f (rel.underS r a)" - proof- - have "\b \ rel.underS r a. inj_on f (rel.under r b)" - using IH by (auto simp add: bij_betw_def) - moreover - have "\b. wo_rel.ofilter r (rel.under r b)" - using Well by (auto simp add: wo_rel.under_ofilter) - ultimately show ?thesis - using WELL bFact UN - UNION_inj_on_ofilter[of r "rel.underS r a" "\b. rel.under r b" f] - by auto - qed - (* *) - have BIJ: "bij_betw f (rel.underS r a) (rel.underS r' (f a))" - unfolding bij_betw_def - using INJ INCL1 INCL2 by auto - (* *) - have "f a \ Field r'" - using Well' subField NE SUC - by (auto simp add: wo_rel.suc_inField) - thus ?thesis - using WELL WELL' IN BIJ under_underS_bij_betw[of r r' a f] by auto -qed - - -lemma wellorders_totally_ordered_aux2: -fixes r ::"'a rel" and r'::"'a' rel" and - f :: "'a \ 'a'" and g :: "'a \ bool" and a::'a -assumes WELL: "Well_order r" and WELL': "Well_order r'" and -MAIN1: - "\ a. (False \ g`(rel.underS r a) \ f`(rel.underS r a) \ Field r' - \ f a = wo_rel.suc r' (f`(rel.underS r a)) \ g a = True) - \ - (\(False \ (g`(rel.underS r a)) \ f`(rel.underS r a) \ Field r') - \ g a = False)" and -MAIN2: "\ a. a \ Field r \ False \ g`(rel.under r a) \ - bij_betw f (rel.under r a) (rel.under r' (f a))" and -Case: "a \ Field r \ False \ g`(rel.under r a)" -shows "\f'. embed r' r f'" -proof- - have Well: "wo_rel r" using WELL unfolding wo_rel_def . - hence Refl: "Refl r" using wo_rel.REFL[of r] by auto - have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto - have Antisym: "antisym r" using Well wo_rel.ANTISYM[of r] by auto - have Well': "wo_rel r'" using WELL' unfolding wo_rel_def . - (* *) - have 0: "rel.under r a = rel.underS r a \ {a}" - using Refl Case by(auto simp add: rel.Refl_under_underS) - (* *) - have 1: "g a = False" - proof- - {assume "g a \ False" - with 0 Case have "False \ g`(rel.underS r a)" by blast - with MAIN1 have "g a = False" by blast} - thus ?thesis by blast - qed - let ?A = "{a \ Field r. g a = False}" - let ?a = "(wo_rel.minim r ?A)" - (* *) - have 2: "?A \ {} \ ?A \ Field r" using Case 1 by blast - (* *) - have 3: "False \ g`(rel.underS r ?a)" - proof - assume "False \ g`(rel.underS r ?a)" - then obtain b where "b \ rel.underS r ?a" and 31: "g b = False" by auto - hence 32: "(b,?a) \ r \ b \ ?a" - by (auto simp add: rel.underS_def) - hence "b \ Field r" unfolding Field_def by auto - with 31 have "b \ ?A" by auto - hence "(?a,b) \ r" using wo_rel.minim_least 2 Well by fastforce - (* again: why worked without type annotations? *) - with 32 Antisym show False - by (auto simp add: antisym_def) - qed - have temp: "?a \ ?A" - using Well 2 wo_rel.minim_in[of r ?A] by auto - hence 4: "?a \ Field r" by auto - (* *) - have 5: "g ?a = False" using temp by blast - (* *) - have 6: "f`(rel.underS r ?a) = Field r'" - using MAIN1[of ?a] 3 5 by blast - (* *) - have 7: "\b \ rel.underS r ?a. bij_betw f (rel.under r b) (rel.under r' (f b))" - proof - fix b assume as: "b \ rel.underS r ?a" - moreover - have "wo_rel.ofilter r (rel.underS r ?a)" - using Well by (auto simp add: wo_rel.underS_ofilter) - ultimately - have "False \ g`(rel.under r b)" using 3 Well by (subst (asm) wo_rel.ofilter_def) fast+ - moreover have "b \ Field r" - unfolding Field_def using as by (auto simp add: rel.underS_def) - ultimately - show "bij_betw f (rel.under r b) (rel.under r' (f b))" - using MAIN2 by auto - qed - (* *) - have "embed r' r (inv_into (rel.underS r ?a) f)" - using WELL WELL' 7 4 6 inv_into_underS_embed[of r ?a f r'] by auto - thus ?thesis - unfolding embed_def by blast -qed - - -theorem wellorders_totally_ordered: -fixes r ::"'a rel" and r'::"'a' rel" -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "(\f. embed r r' f) \ (\f'. embed r' r f')" -proof- - (* Preliminary facts *) - have Well: "wo_rel r" using WELL unfolding wo_rel_def . - hence Refl: "Refl r" using wo_rel.REFL[of r] by auto - have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto - have Well': "wo_rel r'" using WELL' unfolding wo_rel_def . - (* Main proof *) - obtain H where H_def: "H = - (\h a. if False \ (snd o h)`(rel.underS r a) \ (fst o h)`(rel.underS r a) \ Field r' - then (wo_rel.suc r' ((fst o h)`(rel.underS r a)), True) - else (undefined, False))" by blast - have Adm: "wo_rel.adm_wo r H" - using Well - proof(unfold wo_rel.adm_wo_def, clarify) - fix h1::"'a \ 'a' * bool" and h2::"'a \ 'a' * bool" and x - assume "\y\rel.underS r x. h1 y = h2 y" - hence "\y\rel.underS r x. (fst o h1) y = (fst o h2) y \ - (snd o h1) y = (snd o h2) y" by auto - hence "(fst o h1)`(rel.underS r x) = (fst o h2)`(rel.underS r x) \ - (snd o h1)`(rel.underS r x) = (snd o h2)`(rel.underS r x)" - by (auto simp add: image_def) - thus "H h1 x = H h2 x" by (simp add: H_def del: not_False_in_image_Ball) - qed - (* More constant definitions: *) - obtain h::"'a \ 'a' * bool" and f::"'a \ 'a'" and g::"'a \ bool" - where h_def: "h = wo_rel.worec r H" and - f_def: "f = fst o h" and g_def: "g = snd o h" by blast - obtain test where test_def: - "test = (\ a. False \ (g`(rel.underS r a)) \ f`(rel.underS r a) \ Field r')" by blast - (* *) - have *: "\ a. h a = H h a" - using Adm Well wo_rel.worec_fixpoint[of r H] by (simp add: h_def) - have Main1: - "\ a. (test a \ f a = wo_rel.suc r' (f`(rel.underS r a)) \ g a = True) \ - (\(test a) \ g a = False)" - proof- (* How can I prove this withou fixing a? *) - fix a show "(test a \ f a = wo_rel.suc r' (f`(rel.underS r a)) \ g a = True) \ - (\(test a) \ g a = False)" - using *[of a] test_def f_def g_def H_def by auto - qed - (* *) - let ?phi = "\ a. a \ Field r \ False \ g`(rel.under r a) \ - bij_betw f (rel.under r a) (rel.under r' (f a))" - have Main2: "\ a. ?phi a" - proof- - fix a show "?phi a" - proof(rule wo_rel.well_order_induct[of r ?phi], - simp only: Well, clarify) - fix a - assume IH: "\b. b \ a \ (b,a) \ r \ ?phi b" and - *: "a \ Field r" and - **: "False \ g`(rel.under r a)" - have 1: "\b \ rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))" - proof(clarify) - fix b assume ***: "b \ rel.underS r a" - hence 0: "(b,a) \ r \ b \ a" unfolding rel.underS_def by auto - moreover have "b \ Field r" - using *** rel.underS_Field[of r a] by auto - moreover have "False \ g`(rel.under r b)" - using 0 ** Trans rel.under_incr[of r b a] by auto - ultimately show "bij_betw f (rel.under r b) (rel.under r' (f b))" - using IH by auto - qed - (* *) - have 21: "False \ g`(rel.underS r a)" - using ** rel.underS_subset_under[of r a] by auto - have 22: "g`(rel.under r a) \ {True}" using ** by auto - moreover have 23: "a \ rel.under r a" - using Refl * by (auto simp add: rel.Refl_under_in) - ultimately have 24: "g a = True" by blast - have 2: "f`(rel.underS r a) \ Field r'" - proof - assume "f`(rel.underS r a) = Field r'" - hence "g a = False" using Main1 test_def by blast - with 24 show False using ** by blast - qed - (* *) - have 3: "f a = wo_rel.suc r' (f`(rel.underS r a))" - using 21 2 Main1 test_def by blast - (* *) - show "bij_betw f (rel.under r a) (rel.under r' (f a))" - using WELL WELL' 1 2 3 * - wellorders_totally_ordered_aux[of r r' a f] by auto - qed - qed - (* *) - let ?chi = "(\ a. a \ Field r \ False \ g`(rel.under r a))" - show ?thesis - proof(cases "\a. ?chi a") - assume "\ (\a. ?chi a)" - hence "\a \ Field r. bij_betw f (rel.under r a) (rel.under r' (f a))" - using Main2 by blast - thus ?thesis unfolding embed_def by blast - next - assume "\a. ?chi a" - then obtain a where "?chi a" by blast - hence "\f'. embed r' r f'" - using wellorders_totally_ordered_aux2[of r r' g f a] - WELL WELL' Main1 Main2 test_def by blast - thus ?thesis by blast - qed -qed - - -subsection {* Uniqueness of embeddings *} - - -text{* Here we show a fact complementary to the one from the previous subsection -- namely, -that between any two well-orders there is {\em at most} one embedding, and is the one -definable by the expected well-order recursive equation. As a consequence, any two -embeddings of opposite directions are mutually inverse. *} - - -lemma embed_determined: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r r' f" and IN: "a \ Field r" -shows "f a = wo_rel.suc r' (f`(rel.underS r a))" -proof- - have "bij_betw f (rel.underS r a) (rel.underS r' (f a))" - using assms by (auto simp add: embed_underS) - hence "f`(rel.underS r a) = rel.underS r' (f a)" - by (auto simp add: bij_betw_def) - moreover - {have "f a \ Field r'" using IN - using EMB WELL embed_Field[of r r' f] by auto - hence "f a = wo_rel.suc r' (rel.underS r' (f a))" - using WELL' by (auto simp add: wo_rel_def wo_rel.suc_underS) - } - ultimately show ?thesis by simp -qed - - -lemma embed_unique: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMBf: "embed r r' f" and EMBg: "embed r r' g" -shows "a \ Field r \ f a = g a" -proof(rule wo_rel.well_order_induct[of r], auto simp add: WELL wo_rel_def) - fix a - assume IH: "\b. b \ a \ (b,a): r \ b \ Field r \ f b = g b" and - *: "a \ Field r" - hence "\b \ rel.underS r a. f b = g b" - unfolding rel.underS_def by (auto simp add: Field_def) - hence "f`(rel.underS r a) = g`(rel.underS r a)" by force - thus "f a = g a" - using assms * embed_determined[of r r' f a] embed_determined[of r r' g a] by auto -qed - - -lemma embed_bothWays_inverse: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r r' f" and EMB': "embed r' r f'" -shows "(\a \ Field r. f'(f a) = a) \ (\a' \ Field r'. f(f' a') = a')" -proof- - have "embed r r (f' o f)" using assms - by(auto simp add: comp_embed) - moreover have "embed r r id" using assms - by (auto simp add: id_embed) - ultimately have "\a \ Field r. f'(f a) = a" - using assms embed_unique[of r r "f' o f" id] id_def by auto - moreover - {have "embed r' r' (f o f')" using assms - by(auto simp add: comp_embed) - moreover have "embed r' r' id" using assms - by (auto simp add: id_embed) - ultimately have "\a' \ Field r'. f(f' a') = a'" - using assms embed_unique[of r' r' "f o f'" id] id_def by auto - } - ultimately show ?thesis by blast -qed - - -lemma embed_bothWays_bij_betw: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r r' f" and EMB': "embed r' r g" -shows "bij_betw f (Field r) (Field r')" -proof- - let ?A = "Field r" let ?A' = "Field r'" - have "embed r r (g o f) \ embed r' r' (f o g)" - using assms by (auto simp add: comp_embed) - hence 1: "(\a \ ?A. g(f a) = a) \ (\a' \ ?A'. f(g a') = a')" - using WELL id_embed[of r] embed_unique[of r r "g o f" id] - WELL' id_embed[of r'] embed_unique[of r' r' "f o g" id] - id_def by auto - have 2: "(\a \ ?A. f a \ ?A') \ (\a' \ ?A'. g a' \ ?A)" - using assms embed_Field[of r r' f] embed_Field[of r' r g] by blast - (* *) - show ?thesis - proof(unfold bij_betw_def inj_on_def, auto simp add: 2) - fix a b assume *: "a \ ?A" "b \ ?A" and **: "f a = f b" - have "a = g(f a) \ b = g(f b)" using * 1 by auto - with ** show "a = b" by auto - next - fix a' assume *: "a' \ ?A'" - hence "g a' \ ?A \ f(g a') = a'" using 1 2 by auto - thus "a' \ f ` ?A" by force - qed -qed - - -lemma embed_bothWays_iso: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r r' f" and EMB': "embed r' r g" -shows "iso r r' f" -unfolding iso_def using assms by (auto simp add: embed_bothWays_bij_betw) - - -subsection {* More properties of embeddings, strict embeddings and isomorphisms *} - - -lemma embed_bothWays_Field_bij_betw: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and - EMB: "embed r r' f" and EMB': "embed r' r f'" -shows "bij_betw f (Field r) (Field r')" -proof- - have "(\a \ Field r. f'(f a) = a) \ (\a' \ Field r'. f(f' a') = a')" - using assms by (auto simp add: embed_bothWays_inverse) - moreover - have "f`(Field r) \ Field r' \ f' ` (Field r') \ Field r" - using assms by (auto simp add: embed_Field) - ultimately - show ?thesis using bij_betw_byWitness[of "Field r" f' f "Field r'"] by auto -qed - - -lemma embedS_comp_embed: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" - and EMB: "embedS r r' f" and EMB': "embed r' r'' f'" -shows "embedS r r'' (f' o f)" -proof- - let ?g = "(f' o f)" let ?h = "inv_into (Field r) ?g" - have 1: "embed r r' f \ \ (bij_betw f (Field r) (Field r'))" - using EMB by (auto simp add: embedS_def) - hence 2: "embed r r'' ?g" - using WELL EMB' comp_embed[of r r' f r'' f'] by auto - moreover - {assume "bij_betw ?g (Field r) (Field r'')" - hence "embed r'' r ?h" using 2 WELL - by (auto simp add: inv_into_Field_embed_bij_betw) - hence "embed r' r (?h o f')" using WELL' EMB' - by (auto simp add: comp_embed) - hence "bij_betw f (Field r) (Field r')" using WELL WELL' 1 - by (auto simp add: embed_bothWays_Field_bij_betw) - with 1 have False by blast - } - ultimately show ?thesis unfolding embedS_def by auto -qed - - -lemma embed_comp_embedS: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" - and EMB: "embed r r' f" and EMB': "embedS r' r'' f'" -shows "embedS r r'' (f' o f)" -proof- - let ?g = "(f' o f)" let ?h = "inv_into (Field r) ?g" - have 1: "embed r' r'' f' \ \ (bij_betw f' (Field r') (Field r''))" - using EMB' by (auto simp add: embedS_def) - hence 2: "embed r r'' ?g" - using WELL EMB comp_embed[of r r' f r'' f'] by auto - moreover - {assume "bij_betw ?g (Field r) (Field r'')" - hence "embed r'' r ?h" using 2 WELL - by (auto simp add: inv_into_Field_embed_bij_betw) - hence "embed r'' r' (f o ?h)" using WELL'' EMB - by (auto simp add: comp_embed) - hence "bij_betw f' (Field r') (Field r'')" using WELL' WELL'' 1 - by (auto simp add: embed_bothWays_Field_bij_betw) - with 1 have False by blast - } - ultimately show ?thesis unfolding embedS_def by auto -qed - - -lemma embed_comp_iso: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" - and EMB: "embed r r' f" and EMB': "iso r' r'' f'" -shows "embed r r'' (f' o f)" -using assms unfolding iso_def -by (auto simp add: comp_embed) - - -lemma iso_comp_embed: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" - and EMB: "iso r r' f" and EMB': "embed r' r'' f'" -shows "embed r r'' (f' o f)" -using assms unfolding iso_def -by (auto simp add: comp_embed) - - -lemma embedS_comp_iso: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" - and EMB: "embedS r r' f" and EMB': "iso r' r'' f'" -shows "embedS r r'' (f' o f)" -using assms unfolding iso_def -by (auto simp add: embedS_comp_embed) - - -lemma iso_comp_embedS: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" - and EMB: "iso r r' f" and EMB': "embedS r' r'' f'" -shows "embedS r r'' (f' o f)" -using assms unfolding iso_def using embed_comp_embedS -by (auto simp add: embed_comp_embedS) - - -lemma embedS_Field: -assumes WELL: "Well_order r" and EMB: "embedS r r' f" -shows "f ` (Field r) < Field r'" -proof- - have "f`(Field r) \ Field r'" using assms - by (auto simp add: embed_Field embedS_def) - moreover - {have "inj_on f (Field r)" using assms - by (auto simp add: embedS_def embed_inj_on) - hence "f`(Field r) \ Field r'" using EMB - by (auto simp add: embedS_def bij_betw_def) - } - ultimately show ?thesis by blast -qed - - -lemma embedS_iff: -assumes WELL: "Well_order r" and ISO: "embed r r' f" -shows "embedS r r' f = (f ` (Field r) < Field r')" -proof - assume "embedS r r' f" - thus "f ` Field r \ Field r'" - using WELL by (auto simp add: embedS_Field) -next - assume "f ` Field r \ Field r'" - hence "\ bij_betw f (Field r) (Field r')" - unfolding bij_betw_def by blast - thus "embedS r r' f" unfolding embedS_def - using ISO by auto -qed - - -lemma iso_Field: -"iso r r' f \ f ` (Field r) = Field r'" -using assms by (auto simp add: iso_def bij_betw_def) - - -lemma iso_iff: -assumes "Well_order r" -shows "iso r r' f = (embed r r' f \ f ` (Field r) = Field r')" -proof - assume "iso r r' f" - thus "embed r r' f \ f ` (Field r) = Field r'" - by (auto simp add: iso_Field iso_def) -next - assume *: "embed r r' f \ f ` Field r = Field r'" - hence "inj_on f (Field r)" using assms by (auto simp add: embed_inj_on) - with * have "bij_betw f (Field r) (Field r')" - unfolding bij_betw_def by simp - with * show "iso r r' f" unfolding iso_def by auto -qed - - -lemma iso_iff2: -assumes "Well_order r" -shows "iso r r' f = (bij_betw f (Field r) (Field r') \ - (\a \ Field r. \b \ Field r. - (((a,b) \ r) = ((f a, f b) \ r'))))" -using assms -proof(auto simp add: iso_def) - fix a b - assume "embed r r' f" - hence "compat r r' f" using embed_compat[of r] by auto - moreover assume "(a,b) \ r" - ultimately show "(f a, f b) \ r'" using compat_def[of r] by auto -next - let ?f' = "inv_into (Field r) f" - assume "embed r r' f" and 1: "bij_betw f (Field r) (Field r')" - hence "embed r' r ?f'" using assms - by (auto simp add: inv_into_Field_embed_bij_betw) - hence 2: "compat r' r ?f'" using embed_compat[of r'] by auto - fix a b assume *: "a \ Field r" "b \ Field r" and **: "(f a,f b) \ r'" - hence "?f'(f a) = a \ ?f'(f b) = b" using 1 - by (auto simp add: bij_betw_inv_into_left) - thus "(a,b) \ r" using ** 2 compat_def[of r' r ?f'] by fastforce -next - assume *: "bij_betw f (Field r) (Field r')" and - **: "\a\Field r. \b\Field r. ((a, b) \ r) = ((f a, f b) \ r')" - have 1: "\ a. rel.under r a \ Field r \ rel.under r' (f a) \ Field r'" - by (auto simp add: rel.under_Field) - have 2: "inj_on f (Field r)" using * by (auto simp add: bij_betw_def) - {fix a assume ***: "a \ Field r" - have "bij_betw f (rel.under r a) (rel.under r' (f a))" - proof(unfold bij_betw_def, auto) - show "inj_on f (rel.under r a)" - using 1 2 by (auto simp add: subset_inj_on) - next - fix b assume "b \ rel.under r a" - hence "a \ Field r \ b \ Field r \ (b,a) \ r" - unfolding rel.under_def by (auto simp add: Field_def Range_def Domain_def) - with 1 ** show "f b \ rel.under r' (f a)" - unfolding rel.under_def by auto - next - fix b' assume "b' \ rel.under r' (f a)" - hence 3: "(b',f a) \ r'" unfolding rel.under_def by simp - hence "b' \ Field r'" unfolding Field_def by auto - with * obtain b where "b \ Field r \ f b = b'" - unfolding bij_betw_def by force - with 3 ** *** - show "b' \ f ` (rel.under r a)" unfolding rel.under_def by blast - qed - } - thus "embed r r' f" unfolding embed_def using * by auto -qed - - -lemma iso_iff3: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "iso r r' f = (bij_betw f (Field r) (Field r') \ compat r r' f)" -proof - assume "iso r r' f" - thus "bij_betw f (Field r) (Field r') \ compat r r' f" - unfolding compat_def using WELL by (auto simp add: iso_iff2 Field_def) -next - have Well: "wo_rel r \ wo_rel r'" using WELL WELL' - by (auto simp add: wo_rel_def) - assume *: "bij_betw f (Field r) (Field r') \ compat r r' f" - thus "iso r r' f" - unfolding "compat_def" using assms - proof(auto simp add: iso_iff2) - fix a b assume **: "a \ Field r" "b \ Field r" and - ***: "(f a, f b) \ r'" - {assume "(b,a) \ r \ b = a" - hence "(b,a): r"using Well ** wo_rel.REFL[of r] refl_on_def[of _ r] by blast - hence "(f b, f a) \ r'" using * unfolding compat_def by auto - hence "f a = f b" - using Well *** wo_rel.ANTISYM[of r'] antisym_def[of r'] by blast - hence "a = b" using * ** unfolding bij_betw_def inj_on_def by auto - hence "(a,b) \ r" using Well ** wo_rel.REFL[of r] refl_on_def[of _ r] by blast - } - thus "(a,b) \ r" - using Well ** wo_rel.TOTAL[of r] total_on_def[of _ r] by blast - qed -qed - - - -end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellorder_Relation.thy --- a/src/HOL/Cardinals/Wellorder_Relation.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Cardinals/Wellorder_Relation.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* Well-Order Relations *} theory Wellorder_Relation -imports Wellorder_Relation_LFP Wellfounded_More +imports Wellorder_Relation_FP Wellfounded_More begin context wo_rel diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellorder_Relation_FP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Cardinals/Wellorder_Relation_FP.thy Mon Nov 18 18:04:45 2013 +0100 @@ -0,0 +1,631 @@ +(* Title: HOL/Cardinals/Wellorder_Relation_FP.thy + Author: Andrei Popescu, TU Muenchen + Copyright 2012 + +Well-order relations (FP). +*) + +header {* Well-Order Relations (FP) *} + +theory Wellorder_Relation_FP +imports Wellfounded_More_FP +begin + + +text{* In this section, we develop basic concepts and results pertaining +to well-order relations. Note that we consider well-order relations +as {\em non-strict relations}, +i.e., as containing the diagonals of their fields. *} + + +locale wo_rel = rel + assumes WELL: "Well_order r" +begin + +text{* The following context encompasses all this section. In other words, +for the whole section, we consider a fixed well-order relation @{term "r"}. *} + +(* context wo_rel *) + + +subsection {* Auxiliaries *} + + +lemma REFL: "Refl r" +using WELL order_on_defs[of _ r] by auto + + +lemma TRANS: "trans r" +using WELL order_on_defs[of _ r] by auto + + +lemma ANTISYM: "antisym r" +using WELL order_on_defs[of _ r] by auto + + +lemma TOTAL: "Total r" +using WELL order_on_defs[of _ r] by auto + + +lemma TOTALS: "\a \ Field r. \b \ Field r. (a,b) \ r \ (b,a) \ r" +using REFL TOTAL refl_on_def[of _ r] total_on_def[of _ r] by force + + +lemma LIN: "Linear_order r" +using WELL well_order_on_def[of _ r] by auto + + +lemma WF: "wf (r - Id)" +using WELL well_order_on_def[of _ r] by auto + + +lemma cases_Total: +"\ phi a b. \{a,b} <= Field r; ((a,b) \ r \ phi a b); ((b,a) \ r \ phi a b)\ + \ phi a b" +using TOTALS by auto + + +lemma cases_Total3: +"\ phi a b. \{a,b} \ Field r; ((a,b) \ r - Id \ (b,a) \ r - Id \ phi a b); + (a = b \ phi a b)\ \ phi a b" +using TOTALS by auto + + +subsection {* Well-founded induction and recursion adapted to non-strict well-order relations *} + + +text{* Here we provide induction and recursion principles specific to {\em non-strict} +well-order relations. +Although minor variations of those for well-founded relations, they will be useful +for doing away with the tediousness of +having to take out the diagonal each time in order to switch to a well-founded relation. *} + + +lemma well_order_induct: +assumes IND: "\x. \y. y \ x \ (y, x) \ r \ P y \ P x" +shows "P a" +proof- + have "\x. \y. (y, x) \ r - Id \ P y \ P x" + using IND by blast + thus "P a" using WF wf_induct[of "r - Id" P a] by blast +qed + + +definition +worec :: "(('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" +where +"worec F \ wfrec (r - Id) F" + + +definition +adm_wo :: "(('a \ 'b) \ 'a \ 'b) \ bool" +where +"adm_wo H \ \f g x. (\y \ underS x. f y = g y) \ H f x = H g x" + + +lemma worec_fixpoint: +assumes ADM: "adm_wo H" +shows "worec H = H (worec H)" +proof- + let ?rS = "r - Id" + have "adm_wf (r - Id) H" + unfolding adm_wf_def + using ADM adm_wo_def[of H] underS_def by auto + hence "wfrec ?rS H = H (wfrec ?rS H)" + using WF wfrec_fixpoint[of ?rS H] by simp + thus ?thesis unfolding worec_def . +qed + + +subsection {* The notions of maximum, minimum, supremum, successor and order filter *} + + +text{* +We define the successor {\em of a set}, and not of an element (the latter is of course +a particular case). Also, we define the maximum {\em of two elements}, @{text "max2"}, +and the minimum {\em of a set}, @{text "minim"} -- we chose these variants since we +consider them the most useful for well-orders. The minimum is defined in terms of the +auxiliary relational operator @{text "isMinim"}. Then, supremum and successor are +defined in terms of minimum as expected. +The minimum is only meaningful for non-empty sets, and the successor is only +meaningful for sets for which strict upper bounds exist. +Order filters for well-orders are also known as ``initial segments". *} + + +definition max2 :: "'a \ 'a \ 'a" +where "max2 a b \ if (a,b) \ r then b else a" + + +definition isMinim :: "'a set \ 'a \ bool" +where "isMinim A b \ b \ A \ (\a \ A. (b,a) \ r)" + +definition minim :: "'a set \ 'a" +where "minim A \ THE b. isMinim A b" + + +definition supr :: "'a set \ 'a" +where "supr A \ minim (Above A)" + +definition suc :: "'a set \ 'a" +where "suc A \ minim (AboveS A)" + +definition ofilter :: "'a set \ bool" +where +"ofilter A \ (A \ Field r) \ (\a \ A. under a \ A)" + + +subsubsection {* Properties of max2 *} + + +lemma max2_greater_among: +assumes "a \ Field r" and "b \ Field r" +shows "(a, max2 a b) \ r \ (b, max2 a b) \ r \ max2 a b \ {a,b}" +proof- + {assume "(a,b) \ r" + hence ?thesis using max2_def assms REFL refl_on_def + by (auto simp add: refl_on_def) + } + moreover + {assume "a = b" + hence "(a,b) \ r" using REFL assms + by (auto simp add: refl_on_def) + } + moreover + {assume *: "a \ b \ (b,a) \ r" + hence "(a,b) \ r" using ANTISYM + by (auto simp add: antisym_def) + hence ?thesis using * max2_def assms REFL refl_on_def + by (auto simp add: refl_on_def) + } + ultimately show ?thesis using assms TOTAL + total_on_def[of "Field r" r] by blast +qed + + +lemma max2_greater: +assumes "a \ Field r" and "b \ Field r" +shows "(a, max2 a b) \ r \ (b, max2 a b) \ r" +using assms by (auto simp add: max2_greater_among) + + +lemma max2_among: +assumes "a \ Field r" and "b \ Field r" +shows "max2 a b \ {a, b}" +using assms max2_greater_among[of a b] by simp + + +lemma max2_equals1: +assumes "a \ Field r" and "b \ Field r" +shows "(max2 a b = a) = ((b,a) \ r)" +using assms ANTISYM unfolding antisym_def using TOTALS +by(auto simp add: max2_def max2_among) + + +lemma max2_equals2: +assumes "a \ Field r" and "b \ Field r" +shows "(max2 a b = b) = ((a,b) \ r)" +using assms ANTISYM unfolding antisym_def using TOTALS +unfolding max2_def by auto + + +subsubsection {* Existence and uniqueness for isMinim and well-definedness of minim *} + + +lemma isMinim_unique: +assumes MINIM: "isMinim B a" and MINIM': "isMinim B a'" +shows "a = a'" +proof- + {have "a \ B" + using MINIM isMinim_def by simp + hence "(a',a) \ r" + using MINIM' isMinim_def by simp + } + moreover + {have "a' \ B" + using MINIM' isMinim_def by simp + hence "(a,a') \ r" + using MINIM isMinim_def by simp + } + ultimately + show ?thesis using ANTISYM antisym_def[of r] by blast +qed + + +lemma Well_order_isMinim_exists: +assumes SUB: "B \ Field r" and NE: "B \ {}" +shows "\b. isMinim B b" +proof- + from spec[OF WF[unfolded wf_eq_minimal[of "r - Id"]], of B] NE obtain b where + *: "b \ B \ (\b'. b' \ b \ (b',b) \ r \ b' \ B)" by auto + show ?thesis + proof(simp add: isMinim_def, rule exI[of _ b], auto) + show "b \ B" using * by simp + next + fix b' assume As: "b' \ B" + hence **: "b \ Field r \ b' \ Field r" using As SUB * by auto + (* *) + from As * have "b' = b \ (b',b) \ r" by auto + moreover + {assume "b' = b" + hence "(b,b') \ r" + using ** REFL by (auto simp add: refl_on_def) + } + moreover + {assume "b' \ b \ (b',b) \ r" + hence "(b,b') \ r" + using ** TOTAL by (auto simp add: total_on_def) + } + ultimately show "(b,b') \ r" by blast + qed +qed + + +lemma minim_isMinim: +assumes SUB: "B \ Field r" and NE: "B \ {}" +shows "isMinim B (minim B)" +proof- + let ?phi = "(\ b. isMinim B b)" + from assms Well_order_isMinim_exists + obtain b where *: "?phi b" by blast + moreover + have "\ b'. ?phi b' \ b' = b" + using isMinim_unique * by auto + ultimately show ?thesis + unfolding minim_def using theI[of ?phi b] by blast +qed + + +subsubsection{* Properties of minim *} + + +lemma minim_in: +assumes "B \ Field r" and "B \ {}" +shows "minim B \ B" +proof- + from minim_isMinim[of B] assms + have "isMinim B (minim B)" by simp + thus ?thesis by (simp add: isMinim_def) +qed + + +lemma minim_inField: +assumes "B \ Field r" and "B \ {}" +shows "minim B \ Field r" +proof- + have "minim B \ B" using assms by (simp add: minim_in) + thus ?thesis using assms by blast +qed + + +lemma minim_least: +assumes SUB: "B \ Field r" and IN: "b \ B" +shows "(minim B, b) \ r" +proof- + from minim_isMinim[of B] assms + have "isMinim B (minim B)" by auto + thus ?thesis by (auto simp add: isMinim_def IN) +qed + + +lemma equals_minim: +assumes SUB: "B \ Field r" and IN: "a \ B" and + LEAST: "\ b. b \ B \ (a,b) \ r" +shows "a = minim B" +proof- + from minim_isMinim[of B] assms + have "isMinim B (minim B)" by auto + moreover have "isMinim B a" using IN LEAST isMinim_def by auto + ultimately show ?thesis + using isMinim_unique by auto +qed + + +subsubsection{* Properties of successor *} + + +lemma suc_AboveS: +assumes SUB: "B \ Field r" and ABOVES: "AboveS B \ {}" +shows "suc B \ AboveS B" +proof(unfold suc_def) + have "AboveS B \ Field r" + using AboveS_Field by auto + thus "minim (AboveS B) \ AboveS B" + using assms by (simp add: minim_in) +qed + + +lemma suc_greater: +assumes SUB: "B \ Field r" and ABOVES: "AboveS B \ {}" and + IN: "b \ B" +shows "suc B \ b \ (b,suc B) \ r" +proof- + from assms suc_AboveS + have "suc B \ AboveS B" by simp + with IN AboveS_def show ?thesis by simp +qed + + +lemma suc_least_AboveS: +assumes ABOVES: "a \ AboveS B" +shows "(suc B,a) \ r" +proof(unfold suc_def) + have "AboveS B \ Field r" + using AboveS_Field by auto + thus "(minim (AboveS B),a) \ r" + using assms minim_least by simp +qed + + +lemma suc_inField: +assumes "B \ Field r" and "AboveS B \ {}" +shows "suc B \ Field r" +proof- + have "suc B \ AboveS B" using suc_AboveS assms by simp + thus ?thesis + using assms AboveS_Field by auto +qed + + +lemma equals_suc_AboveS: +assumes SUB: "B \ Field r" and ABV: "a \ AboveS B" and + MINIM: "\ a'. a' \ AboveS B \ (a,a') \ r" +shows "a = suc B" +proof(unfold suc_def) + have "AboveS B \ Field r" + using AboveS_Field[of B] by auto + thus "a = minim (AboveS B)" + using assms equals_minim + by simp +qed + + +lemma suc_underS: +assumes IN: "a \ Field r" +shows "a = suc (underS a)" +proof- + have "underS a \ Field r" + using underS_Field by auto + moreover + have "a \ AboveS (underS a)" + using in_AboveS_underS IN by auto + moreover + have "\a' \ AboveS (underS a). (a,a') \ r" + proof(clarify) + fix a' + assume *: "a' \ AboveS (underS a)" + hence **: "a' \ Field r" + using AboveS_Field by auto + {assume "(a,a') \ r" + hence "a' = a \ (a',a) \ r" + using TOTAL IN ** by (auto simp add: total_on_def) + moreover + {assume "a' = a" + hence "(a,a') \ r" + using REFL IN ** by (auto simp add: refl_on_def) + } + moreover + {assume "a' \ a \ (a',a) \ r" + hence "a' \ underS a" + unfolding underS_def by simp + hence "a' \ AboveS (underS a)" + using AboveS_disjoint by blast + with * have False by simp + } + ultimately have "(a,a') \ r" by blast + } + thus "(a, a') \ r" by blast + qed + ultimately show ?thesis + using equals_suc_AboveS by auto +qed + + +subsubsection {* Properties of order filters *} + + +lemma under_ofilter: +"ofilter (under a)" +proof(unfold ofilter_def under_def, auto simp add: Field_def) + fix aa x + assume "(aa,a) \ r" "(x,aa) \ r" + thus "(x,a) \ r" + using TRANS trans_def[of r] by blast +qed + + +lemma underS_ofilter: +"ofilter (underS a)" +proof(unfold ofilter_def underS_def under_def, auto simp add: Field_def) + fix aa assume "(a, aa) \ r" "(aa, a) \ r" and DIFF: "aa \ a" + thus False + using ANTISYM antisym_def[of r] by blast +next + fix aa x + assume "(aa,a) \ r" "aa \ a" "(x,aa) \ r" + thus "(x,a) \ r" + using TRANS trans_def[of r] by blast +qed + + +lemma Field_ofilter: +"ofilter (Field r)" +by(unfold ofilter_def under_def, auto simp add: Field_def) + + +lemma ofilter_underS_Field: +"ofilter A = ((\a \ Field r. A = underS a) \ (A = Field r))" +proof + assume "(\a\Field r. A = underS a) \ A = Field r" + thus "ofilter A" + by (auto simp: underS_ofilter Field_ofilter) +next + assume *: "ofilter A" + let ?One = "(\a\Field r. A = underS a)" + let ?Two = "(A = Field r)" + show "?One \ ?Two" + proof(cases ?Two, simp) + let ?B = "(Field r) - A" + let ?a = "minim ?B" + assume "A \ Field r" + moreover have "A \ Field r" using * ofilter_def by simp + ultimately have 1: "?B \ {}" by blast + hence 2: "?a \ Field r" using minim_inField[of ?B] by blast + have 3: "?a \ ?B" using minim_in[of ?B] 1 by blast + hence 4: "?a \ A" by blast + have 5: "A \ Field r" using * ofilter_def[of A] by auto + (* *) + moreover + have "A = underS ?a" + proof + show "A \ underS ?a" + proof(unfold underS_def, auto simp add: 4) + fix x assume **: "x \ A" + hence 11: "x \ Field r" using 5 by auto + have 12: "x \ ?a" using 4 ** by auto + have 13: "under x \ A" using * ofilter_def ** by auto + {assume "(x,?a) \ r" + hence "(?a,x) \ r" + using TOTAL total_on_def[of "Field r" r] + 2 4 11 12 by auto + hence "?a \ under x" using under_def by auto + hence "?a \ A" using ** 13 by blast + with 4 have False by simp + } + thus "(x,?a) \ r" by blast + qed + next + show "underS ?a \ A" + proof(unfold underS_def, auto) + fix x + assume **: "x \ ?a" and ***: "(x,?a) \ r" + hence 11: "x \ Field r" using Field_def by fastforce + {assume "x \ A" + hence "x \ ?B" using 11 by auto + hence "(?a,x) \ r" using 3 minim_least[of ?B x] by blast + hence False + using ANTISYM antisym_def[of r] ** *** by auto + } + thus "x \ A" by blast + qed + qed + ultimately have ?One using 2 by blast + thus ?thesis by simp + qed +qed + + +lemma ofilter_UNION: +"(\ i. i \ I \ ofilter(A i)) \ ofilter (\ i \ I. A i)" +unfolding ofilter_def by blast + + +lemma ofilter_under_UNION: +assumes "ofilter A" +shows "A = (\ a \ A. under a)" +proof + have "\a \ A. under a \ A" + using assms ofilter_def by auto + thus "(\ a \ A. under a) \ A" by blast +next + have "\a \ A. a \ under a" + using REFL Refl_under_in assms ofilter_def by blast + thus "A \ (\ a \ A. under a)" by blast +qed + + +subsubsection{* Other properties *} + + +lemma ofilter_linord: +assumes OF1: "ofilter A" and OF2: "ofilter B" +shows "A \ B \ B \ A" +proof(cases "A = Field r") + assume Case1: "A = Field r" + hence "B \ A" using OF2 ofilter_def by auto + thus ?thesis by simp +next + assume Case2: "A \ Field r" + with ofilter_underS_Field OF1 obtain a where + 1: "a \ Field r \ A = underS a" by auto + show ?thesis + proof(cases "B = Field r") + assume Case21: "B = Field r" + hence "A \ B" using OF1 ofilter_def by auto + thus ?thesis by simp + next + assume Case22: "B \ Field r" + with ofilter_underS_Field OF2 obtain b where + 2: "b \ Field r \ B = underS b" by auto + have "a = b \ (a,b) \ r \ (b,a) \ r" + using 1 2 TOTAL total_on_def[of _ r] by auto + moreover + {assume "a = b" with 1 2 have ?thesis by auto + } + moreover + {assume "(a,b) \ r" + with underS_incr TRANS ANTISYM 1 2 + have "A \ B" by auto + hence ?thesis by auto + } + moreover + {assume "(b,a) \ r" + with underS_incr TRANS ANTISYM 1 2 + have "B \ A" by auto + hence ?thesis by auto + } + ultimately show ?thesis by blast + qed +qed + + +lemma ofilter_AboveS_Field: +assumes "ofilter A" +shows "A \ (AboveS A) = Field r" +proof + show "A \ (AboveS A) \ Field r" + using assms ofilter_def AboveS_Field by auto +next + {fix x assume *: "x \ Field r" and **: "x \ A" + {fix y assume ***: "y \ A" + with ** have 1: "y \ x" by auto + {assume "(y,x) \ r" + moreover + have "y \ Field r" using assms ofilter_def *** by auto + ultimately have "(x,y) \ r" + using 1 * TOTAL total_on_def[of _ r] by auto + with *** assms ofilter_def under_def have "x \ A" by auto + with ** have False by contradiction + } + hence "(y,x) \ r" by blast + with 1 have "y \ x \ (y,x) \ r" by auto + } + with * have "x \ AboveS A" unfolding AboveS_def by auto + } + thus "Field r \ A \ (AboveS A)" by blast +qed + + +lemma suc_ofilter_in: +assumes OF: "ofilter A" and ABOVE_NE: "AboveS A \ {}" and + REL: "(b,suc A) \ r" and DIFF: "b \ suc A" +shows "b \ A" +proof- + have *: "suc A \ Field r \ b \ Field r" + using WELL REL well_order_on_domain by auto + {assume **: "b \ A" + hence "b \ AboveS A" + using OF * ofilter_AboveS_Field by auto + hence "(suc A, b) \ r" + using suc_least_AboveS by auto + hence False using REL DIFF ANTISYM * + by (auto simp add: antisym_def) + } + thus ?thesis by blast +qed + + + +end (* context wo_rel *) + + + +end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Cardinals/Wellorder_Relation_LFP.thy --- a/src/HOL/Cardinals/Wellorder_Relation_LFP.thy Mon Nov 18 18:04:45 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,631 +0,0 @@ -(* Title: HOL/Cardinals/Wellorder_Relation_LFP.thy - Author: Andrei Popescu, TU Muenchen - Copyright 2012 - -Well-order relations (LFP). -*) - -header {* Well-Order Relations (LFP) *} - -theory Wellorder_Relation_LFP -imports Wellfounded_More_LFP -begin - - -text{* In this section, we develop basic concepts and results pertaining -to well-order relations. Note that we consider well-order relations -as {\em non-strict relations}, -i.e., as containing the diagonals of their fields. *} - - -locale wo_rel = rel + assumes WELL: "Well_order r" -begin - -text{* The following context encompasses all this section. In other words, -for the whole section, we consider a fixed well-order relation @{term "r"}. *} - -(* context wo_rel *) - - -subsection {* Auxiliaries *} - - -lemma REFL: "Refl r" -using WELL order_on_defs[of _ r] by auto - - -lemma TRANS: "trans r" -using WELL order_on_defs[of _ r] by auto - - -lemma ANTISYM: "antisym r" -using WELL order_on_defs[of _ r] by auto - - -lemma TOTAL: "Total r" -using WELL order_on_defs[of _ r] by auto - - -lemma TOTALS: "\a \ Field r. \b \ Field r. (a,b) \ r \ (b,a) \ r" -using REFL TOTAL refl_on_def[of _ r] total_on_def[of _ r] by force - - -lemma LIN: "Linear_order r" -using WELL well_order_on_def[of _ r] by auto - - -lemma WF: "wf (r - Id)" -using WELL well_order_on_def[of _ r] by auto - - -lemma cases_Total: -"\ phi a b. \{a,b} <= Field r; ((a,b) \ r \ phi a b); ((b,a) \ r \ phi a b)\ - \ phi a b" -using TOTALS by auto - - -lemma cases_Total3: -"\ phi a b. \{a,b} \ Field r; ((a,b) \ r - Id \ (b,a) \ r - Id \ phi a b); - (a = b \ phi a b)\ \ phi a b" -using TOTALS by auto - - -subsection {* Well-founded induction and recursion adapted to non-strict well-order relations *} - - -text{* Here we provide induction and recursion principles specific to {\em non-strict} -well-order relations. -Although minor variations of those for well-founded relations, they will be useful -for doing away with the tediousness of -having to take out the diagonal each time in order to switch to a well-founded relation. *} - - -lemma well_order_induct: -assumes IND: "\x. \y. y \ x \ (y, x) \ r \ P y \ P x" -shows "P a" -proof- - have "\x. \y. (y, x) \ r - Id \ P y \ P x" - using IND by blast - thus "P a" using WF wf_induct[of "r - Id" P a] by blast -qed - - -definition -worec :: "(('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" -where -"worec F \ wfrec (r - Id) F" - - -definition -adm_wo :: "(('a \ 'b) \ 'a \ 'b) \ bool" -where -"adm_wo H \ \f g x. (\y \ underS x. f y = g y) \ H f x = H g x" - - -lemma worec_fixpoint: -assumes ADM: "adm_wo H" -shows "worec H = H (worec H)" -proof- - let ?rS = "r - Id" - have "adm_wf (r - Id) H" - unfolding adm_wf_def - using ADM adm_wo_def[of H] underS_def by auto - hence "wfrec ?rS H = H (wfrec ?rS H)" - using WF wfrec_fixpoint[of ?rS H] by simp - thus ?thesis unfolding worec_def . -qed - - -subsection {* The notions of maximum, minimum, supremum, successor and order filter *} - - -text{* -We define the successor {\em of a set}, and not of an element (the latter is of course -a particular case). Also, we define the maximum {\em of two elements}, @{text "max2"}, -and the minimum {\em of a set}, @{text "minim"} -- we chose these variants since we -consider them the most useful for well-orders. The minimum is defined in terms of the -auxiliary relational operator @{text "isMinim"}. Then, supremum and successor are -defined in terms of minimum as expected. -The minimum is only meaningful for non-empty sets, and the successor is only -meaningful for sets for which strict upper bounds exist. -Order filters for well-orders are also known as ``initial segments". *} - - -definition max2 :: "'a \ 'a \ 'a" -where "max2 a b \ if (a,b) \ r then b else a" - - -definition isMinim :: "'a set \ 'a \ bool" -where "isMinim A b \ b \ A \ (\a \ A. (b,a) \ r)" - -definition minim :: "'a set \ 'a" -where "minim A \ THE b. isMinim A b" - - -definition supr :: "'a set \ 'a" -where "supr A \ minim (Above A)" - -definition suc :: "'a set \ 'a" -where "suc A \ minim (AboveS A)" - -definition ofilter :: "'a set \ bool" -where -"ofilter A \ (A \ Field r) \ (\a \ A. under a \ A)" - - -subsubsection {* Properties of max2 *} - - -lemma max2_greater_among: -assumes "a \ Field r" and "b \ Field r" -shows "(a, max2 a b) \ r \ (b, max2 a b) \ r \ max2 a b \ {a,b}" -proof- - {assume "(a,b) \ r" - hence ?thesis using max2_def assms REFL refl_on_def - by (auto simp add: refl_on_def) - } - moreover - {assume "a = b" - hence "(a,b) \ r" using REFL assms - by (auto simp add: refl_on_def) - } - moreover - {assume *: "a \ b \ (b,a) \ r" - hence "(a,b) \ r" using ANTISYM - by (auto simp add: antisym_def) - hence ?thesis using * max2_def assms REFL refl_on_def - by (auto simp add: refl_on_def) - } - ultimately show ?thesis using assms TOTAL - total_on_def[of "Field r" r] by blast -qed - - -lemma max2_greater: -assumes "a \ Field r" and "b \ Field r" -shows "(a, max2 a b) \ r \ (b, max2 a b) \ r" -using assms by (auto simp add: max2_greater_among) - - -lemma max2_among: -assumes "a \ Field r" and "b \ Field r" -shows "max2 a b \ {a, b}" -using assms max2_greater_among[of a b] by simp - - -lemma max2_equals1: -assumes "a \ Field r" and "b \ Field r" -shows "(max2 a b = a) = ((b,a) \ r)" -using assms ANTISYM unfolding antisym_def using TOTALS -by(auto simp add: max2_def max2_among) - - -lemma max2_equals2: -assumes "a \ Field r" and "b \ Field r" -shows "(max2 a b = b) = ((a,b) \ r)" -using assms ANTISYM unfolding antisym_def using TOTALS -unfolding max2_def by auto - - -subsubsection {* Existence and uniqueness for isMinim and well-definedness of minim *} - - -lemma isMinim_unique: -assumes MINIM: "isMinim B a" and MINIM': "isMinim B a'" -shows "a = a'" -proof- - {have "a \ B" - using MINIM isMinim_def by simp - hence "(a',a) \ r" - using MINIM' isMinim_def by simp - } - moreover - {have "a' \ B" - using MINIM' isMinim_def by simp - hence "(a,a') \ r" - using MINIM isMinim_def by simp - } - ultimately - show ?thesis using ANTISYM antisym_def[of r] by blast -qed - - -lemma Well_order_isMinim_exists: -assumes SUB: "B \ Field r" and NE: "B \ {}" -shows "\b. isMinim B b" -proof- - from spec[OF WF[unfolded wf_eq_minimal[of "r - Id"]], of B] NE obtain b where - *: "b \ B \ (\b'. b' \ b \ (b',b) \ r \ b' \ B)" by auto - show ?thesis - proof(simp add: isMinim_def, rule exI[of _ b], auto) - show "b \ B" using * by simp - next - fix b' assume As: "b' \ B" - hence **: "b \ Field r \ b' \ Field r" using As SUB * by auto - (* *) - from As * have "b' = b \ (b',b) \ r" by auto - moreover - {assume "b' = b" - hence "(b,b') \ r" - using ** REFL by (auto simp add: refl_on_def) - } - moreover - {assume "b' \ b \ (b',b) \ r" - hence "(b,b') \ r" - using ** TOTAL by (auto simp add: total_on_def) - } - ultimately show "(b,b') \ r" by blast - qed -qed - - -lemma minim_isMinim: -assumes SUB: "B \ Field r" and NE: "B \ {}" -shows "isMinim B (minim B)" -proof- - let ?phi = "(\ b. isMinim B b)" - from assms Well_order_isMinim_exists - obtain b where *: "?phi b" by blast - moreover - have "\ b'. ?phi b' \ b' = b" - using isMinim_unique * by auto - ultimately show ?thesis - unfolding minim_def using theI[of ?phi b] by blast -qed - - -subsubsection{* Properties of minim *} - - -lemma minim_in: -assumes "B \ Field r" and "B \ {}" -shows "minim B \ B" -proof- - from minim_isMinim[of B] assms - have "isMinim B (minim B)" by simp - thus ?thesis by (simp add: isMinim_def) -qed - - -lemma minim_inField: -assumes "B \ Field r" and "B \ {}" -shows "minim B \ Field r" -proof- - have "minim B \ B" using assms by (simp add: minim_in) - thus ?thesis using assms by blast -qed - - -lemma minim_least: -assumes SUB: "B \ Field r" and IN: "b \ B" -shows "(minim B, b) \ r" -proof- - from minim_isMinim[of B] assms - have "isMinim B (minim B)" by auto - thus ?thesis by (auto simp add: isMinim_def IN) -qed - - -lemma equals_minim: -assumes SUB: "B \ Field r" and IN: "a \ B" and - LEAST: "\ b. b \ B \ (a,b) \ r" -shows "a = minim B" -proof- - from minim_isMinim[of B] assms - have "isMinim B (minim B)" by auto - moreover have "isMinim B a" using IN LEAST isMinim_def by auto - ultimately show ?thesis - using isMinim_unique by auto -qed - - -subsubsection{* Properties of successor *} - - -lemma suc_AboveS: -assumes SUB: "B \ Field r" and ABOVES: "AboveS B \ {}" -shows "suc B \ AboveS B" -proof(unfold suc_def) - have "AboveS B \ Field r" - using AboveS_Field by auto - thus "minim (AboveS B) \ AboveS B" - using assms by (simp add: minim_in) -qed - - -lemma suc_greater: -assumes SUB: "B \ Field r" and ABOVES: "AboveS B \ {}" and - IN: "b \ B" -shows "suc B \ b \ (b,suc B) \ r" -proof- - from assms suc_AboveS - have "suc B \ AboveS B" by simp - with IN AboveS_def show ?thesis by simp -qed - - -lemma suc_least_AboveS: -assumes ABOVES: "a \ AboveS B" -shows "(suc B,a) \ r" -proof(unfold suc_def) - have "AboveS B \ Field r" - using AboveS_Field by auto - thus "(minim (AboveS B),a) \ r" - using assms minim_least by simp -qed - - -lemma suc_inField: -assumes "B \ Field r" and "AboveS B \ {}" -shows "suc B \ Field r" -proof- - have "suc B \ AboveS B" using suc_AboveS assms by simp - thus ?thesis - using assms AboveS_Field by auto -qed - - -lemma equals_suc_AboveS: -assumes SUB: "B \ Field r" and ABV: "a \ AboveS B" and - MINIM: "\ a'. a' \ AboveS B \ (a,a') \ r" -shows "a = suc B" -proof(unfold suc_def) - have "AboveS B \ Field r" - using AboveS_Field[of B] by auto - thus "a = minim (AboveS B)" - using assms equals_minim - by simp -qed - - -lemma suc_underS: -assumes IN: "a \ Field r" -shows "a = suc (underS a)" -proof- - have "underS a \ Field r" - using underS_Field by auto - moreover - have "a \ AboveS (underS a)" - using in_AboveS_underS IN by auto - moreover - have "\a' \ AboveS (underS a). (a,a') \ r" - proof(clarify) - fix a' - assume *: "a' \ AboveS (underS a)" - hence **: "a' \ Field r" - using AboveS_Field by auto - {assume "(a,a') \ r" - hence "a' = a \ (a',a) \ r" - using TOTAL IN ** by (auto simp add: total_on_def) - moreover - {assume "a' = a" - hence "(a,a') \ r" - using REFL IN ** by (auto simp add: refl_on_def) - } - moreover - {assume "a' \ a \ (a',a) \ r" - hence "a' \ underS a" - unfolding underS_def by simp - hence "a' \ AboveS (underS a)" - using AboveS_disjoint by blast - with * have False by simp - } - ultimately have "(a,a') \ r" by blast - } - thus "(a, a') \ r" by blast - qed - ultimately show ?thesis - using equals_suc_AboveS by auto -qed - - -subsubsection {* Properties of order filters *} - - -lemma under_ofilter: -"ofilter (under a)" -proof(unfold ofilter_def under_def, auto simp add: Field_def) - fix aa x - assume "(aa,a) \ r" "(x,aa) \ r" - thus "(x,a) \ r" - using TRANS trans_def[of r] by blast -qed - - -lemma underS_ofilter: -"ofilter (underS a)" -proof(unfold ofilter_def underS_def under_def, auto simp add: Field_def) - fix aa assume "(a, aa) \ r" "(aa, a) \ r" and DIFF: "aa \ a" - thus False - using ANTISYM antisym_def[of r] by blast -next - fix aa x - assume "(aa,a) \ r" "aa \ a" "(x,aa) \ r" - thus "(x,a) \ r" - using TRANS trans_def[of r] by blast -qed - - -lemma Field_ofilter: -"ofilter (Field r)" -by(unfold ofilter_def under_def, auto simp add: Field_def) - - -lemma ofilter_underS_Field: -"ofilter A = ((\a \ Field r. A = underS a) \ (A = Field r))" -proof - assume "(\a\Field r. A = underS a) \ A = Field r" - thus "ofilter A" - by (auto simp: underS_ofilter Field_ofilter) -next - assume *: "ofilter A" - let ?One = "(\a\Field r. A = underS a)" - let ?Two = "(A = Field r)" - show "?One \ ?Two" - proof(cases ?Two, simp) - let ?B = "(Field r) - A" - let ?a = "minim ?B" - assume "A \ Field r" - moreover have "A \ Field r" using * ofilter_def by simp - ultimately have 1: "?B \ {}" by blast - hence 2: "?a \ Field r" using minim_inField[of ?B] by blast - have 3: "?a \ ?B" using minim_in[of ?B] 1 by blast - hence 4: "?a \ A" by blast - have 5: "A \ Field r" using * ofilter_def[of A] by auto - (* *) - moreover - have "A = underS ?a" - proof - show "A \ underS ?a" - proof(unfold underS_def, auto simp add: 4) - fix x assume **: "x \ A" - hence 11: "x \ Field r" using 5 by auto - have 12: "x \ ?a" using 4 ** by auto - have 13: "under x \ A" using * ofilter_def ** by auto - {assume "(x,?a) \ r" - hence "(?a,x) \ r" - using TOTAL total_on_def[of "Field r" r] - 2 4 11 12 by auto - hence "?a \ under x" using under_def by auto - hence "?a \ A" using ** 13 by blast - with 4 have False by simp - } - thus "(x,?a) \ r" by blast - qed - next - show "underS ?a \ A" - proof(unfold underS_def, auto) - fix x - assume **: "x \ ?a" and ***: "(x,?a) \ r" - hence 11: "x \ Field r" using Field_def by fastforce - {assume "x \ A" - hence "x \ ?B" using 11 by auto - hence "(?a,x) \ r" using 3 minim_least[of ?B x] by blast - hence False - using ANTISYM antisym_def[of r] ** *** by auto - } - thus "x \ A" by blast - qed - qed - ultimately have ?One using 2 by blast - thus ?thesis by simp - qed -qed - - -lemma ofilter_UNION: -"(\ i. i \ I \ ofilter(A i)) \ ofilter (\ i \ I. A i)" -unfolding ofilter_def by blast - - -lemma ofilter_under_UNION: -assumes "ofilter A" -shows "A = (\ a \ A. under a)" -proof - have "\a \ A. under a \ A" - using assms ofilter_def by auto - thus "(\ a \ A. under a) \ A" by blast -next - have "\a \ A. a \ under a" - using REFL Refl_under_in assms ofilter_def by blast - thus "A \ (\ a \ A. under a)" by blast -qed - - -subsubsection{* Other properties *} - - -lemma ofilter_linord: -assumes OF1: "ofilter A" and OF2: "ofilter B" -shows "A \ B \ B \ A" -proof(cases "A = Field r") - assume Case1: "A = Field r" - hence "B \ A" using OF2 ofilter_def by auto - thus ?thesis by simp -next - assume Case2: "A \ Field r" - with ofilter_underS_Field OF1 obtain a where - 1: "a \ Field r \ A = underS a" by auto - show ?thesis - proof(cases "B = Field r") - assume Case21: "B = Field r" - hence "A \ B" using OF1 ofilter_def by auto - thus ?thesis by simp - next - assume Case22: "B \ Field r" - with ofilter_underS_Field OF2 obtain b where - 2: "b \ Field r \ B = underS b" by auto - have "a = b \ (a,b) \ r \ (b,a) \ r" - using 1 2 TOTAL total_on_def[of _ r] by auto - moreover - {assume "a = b" with 1 2 have ?thesis by auto - } - moreover - {assume "(a,b) \ r" - with underS_incr TRANS ANTISYM 1 2 - have "A \ B" by auto - hence ?thesis by auto - } - moreover - {assume "(b,a) \ r" - with underS_incr TRANS ANTISYM 1 2 - have "B \ A" by auto - hence ?thesis by auto - } - ultimately show ?thesis by blast - qed -qed - - -lemma ofilter_AboveS_Field: -assumes "ofilter A" -shows "A \ (AboveS A) = Field r" -proof - show "A \ (AboveS A) \ Field r" - using assms ofilter_def AboveS_Field by auto -next - {fix x assume *: "x \ Field r" and **: "x \ A" - {fix y assume ***: "y \ A" - with ** have 1: "y \ x" by auto - {assume "(y,x) \ r" - moreover - have "y \ Field r" using assms ofilter_def *** by auto - ultimately have "(x,y) \ r" - using 1 * TOTAL total_on_def[of _ r] by auto - with *** assms ofilter_def under_def have "x \ A" by auto - with ** have False by contradiction - } - hence "(y,x) \ r" by blast - with 1 have "y \ x \ (y,x) \ r" by auto - } - with * have "x \ AboveS A" unfolding AboveS_def by auto - } - thus "Field r \ A \ (AboveS A)" by blast -qed - - -lemma suc_ofilter_in: -assumes OF: "ofilter A" and ABOVE_NE: "AboveS A \ {}" and - REL: "(b,suc A) \ r" and DIFF: "b \ suc A" -shows "b \ A" -proof- - have *: "suc A \ Field r \ b \ Field r" - using WELL REL well_order_on_domain by auto - {assume **: "b \ A" - hence "b \ AboveS A" - using OF * ofilter_AboveS_Field by auto - hence "(suc A, b) \ r" - using suc_least_AboveS by auto - hence False using REL DIFF ANTISYM * - by (auto simp add: antisym_def) - } - thus ?thesis by blast -qed - - - -end (* context wo_rel *) - - - -end diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/Library/Order_Union.thy --- a/src/HOL/Library/Order_Union.thy Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/Library/Order_Union.thy Mon Nov 18 18:04:45 2013 +0100 @@ -7,7 +7,7 @@ header {* Order Union *} theory Order_Union -imports "~~/src/HOL/Cardinals/Wellfounded_More_LFP" +imports "~~/src/HOL/Cardinals/Wellfounded_More_FP" begin definition Osum :: "'a rel \ 'a rel \ 'a rel" (infix "Osum" 60) where diff -r 57e781b711b5 -r 5c9819d7713b src/HOL/ROOT --- a/src/HOL/ROOT Mon Nov 18 18:04:45 2013 +0100 +++ b/src/HOL/ROOT Mon Nov 18 18:04:45 2013 +0100 @@ -687,21 +687,14 @@ theories Nominal_Examples theories [quick_and_dirty] VC_Condition -session "HOL-Cardinals-LFP" in Cardinals = HOL + +session "HOL-Cardinals-FP" in Cardinals = HOL + description {* - Ordinals and Cardinals, Theories Needed for BNF LFP Construction. + Ordinals and Cardinals, Theories Needed for BNF FP Constructions. *} options [document = false] - theories Cardinal_Arithmetic_LFP + theories Cardinal_Arithmetic_FP -session "HOL-Cardinals-GFP" in Cardinals = "HOL-Cardinals-LFP" + - description {* - Ordinals and Cardinals, Theories Needed for BNF GFP Construction. - *} - options [document = false] - theories Cardinal_Arithmetic_GFP - -session "HOL-Cardinals" in Cardinals = "HOL-Cardinals-GFP" + +session "HOL-Cardinals" in Cardinals = "HOL-Cardinals-FP" + description {* Ordinals and Cardinals, Full Theories. *} @@ -712,23 +705,16 @@ "document/root.tex" "document/root.bib" -session "HOL-BNF-LFP" in BNF = "HOL-Cardinals-LFP" + +session "HOL-BNF-FP" in BNF = "HOL-Cardinals-FP" + description {* - Bounded Natural Functors for Datatypes. + Bounded Natural Functors for (Co)datatypes. *} options [document = false] - theories BNF_LFP + theories BNF_LFP BNF_GFP -session "HOL-BNF-GFP" in BNF = "HOL-Cardinals-GFP" + +session "HOL-BNF" in BNF = "HOL-BNF-FP" + description {* - Bounded Natural Functors for Codatatypes. - *} - options [document = false] - theories BNF_GFP - -session "HOL-BNF" in BNF = "HOL-Cardinals-GFP" + - description {* - Bounded Natural Functors for (Co)datatypes. + Bounded Natural Functors for (Co)datatypes, Including More BNFs. *} options [document = false] theories BNF