# HG changeset patch # User wenzelm # Date 1673809244 -3600 # Node ID d3de24c50b08d78589fbaddfe8d46e1b8adff1cd # Parent f70d431b5016de2a6487c88b908573cb9a0162d4# Parent f327ae3cab2afc5524d4f253a582aad425765a3d merged diff -r f327ae3cab2a -r d3de24c50b08 NEWS --- a/NEWS Sun Jan 15 20:00:37 2023 +0100 +++ b/NEWS Sun Jan 15 20:00:44 2023 +0100 @@ -199,6 +199,9 @@ multp_mono_strong wfP_subset_mset[simp] +* Mirabelle: + - Added session to output directory structure. Minor INCOMPATIBILITY. + *** ML *** diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Analysis/Further_Topology.thy Sun Jan 15 20:00:44 2023 +0100 @@ -168,8 +168,7 @@ by (simp add: \subspace S\ closed_subspace compact_Int_closed) then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \ S) \ T" and g12: "\x. x \ sphere 0 1 \ S \ norm(f x - g x) < 1/2" - apply (rule Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \subspace T\, of "1/2"]) - using fim by auto + using Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \subspace T\, of "1/2"] fim by auto have gnz: "g x \ 0" if "x \ sphere 0 1 \ S" for x proof - have "norm (f x) = 1" @@ -326,8 +325,8 @@ obtain T':: "'a set" where "subspace T'" and affT': "aff_dim T' = aff_dim T" and homT: "rel_frontier T homeomorphic sphere 0 1 \ T'" - apply (rule spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a]) - using \T \ {}\ by (auto simp add: aff_dim_le_DIM) + using \T \ {}\ spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a] + by (force simp add: aff_dim_le_DIM) with homeomorphic_imp_homotopy_eqv have relT: "sphere 0 1 \ T' homotopy_eqv rel_frontier T" using homotopy_equivalent_space_sym by blast @@ -344,10 +343,9 @@ have dimST': "dim S' < dim T'" by (metis \S' \ T'\ \subspace S'\ \subspace T'\ affS' affST affT' less_irrefl not_le subspace_dim_equal) have "\c. homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" - apply (rule homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim]) - apply (rule homotopy_eqv_cohomotopic_triviality_null[OF relS, THEN iffD1, rule_format]) - apply (metis dimST' \subspace S'\ \subspace T'\ \S' \ T'\ spheremap_lemma2, blast) - done + using spheremap_lemma2 homotopy_eqv_cohomotopic_triviality_null[OF relS] + using homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim] + by (metis \S' \ T'\ \subspace S'\ \subspace T'\ dimST') with that show ?thesis by blast qed qed @@ -370,8 +368,8 @@ case False with \\ s \ 0\ have "r > 0" "s > 0" by auto show thesis - apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) - using \0 < r\ \0 < s\ assms(1) that by (simp_all add: f aff_dim_cball) + using inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f] + using \0 < r\ \0 < s\ assms(1) that by (auto simp add: f aff_dim_cball) qed qed @@ -401,12 +399,10 @@ then show ?thesis using UnionI feq geq \S \ \\ subsetD that by fastforce qed - show ?case - apply (rule_tac x="\x. if x \ S then f x else g x" in exI, simp) - apply (intro conjI continuous_on_cases) - using fim gim feq geq - apply (force simp: insert closed_Union contf contg inf_commute intro: fg)+ - done + moreover have "continuous_on (S \ \ \) (\x. if x \ S then f x else g x)" + by (auto simp: insert closed_Union contf contg intro: fg continuous_on_cases) + ultimately show ?case + by (smt (verit, del_insts) Int_iff UnE complete_lattice_class.Sup_insert feq fim geq gim image_subset_iff) qed lemma extending_maps_Union: @@ -573,13 +569,10 @@ have ff: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X face_of D" "Y face_of E" and DE: "D \ \" "E \ \" for D E by (rule face_of_Int_subface [OF _ _ XY]) (auto simp: face' DE) - show ?thesis - using that - apply auto - apply (drule_tac x="X \ Y" in spec, safe) - using ff face_of_imp_convex [of X] face_of_imp_convex [of Y] - apply (fastforce dest: face_of_aff_dim_lt) - by (meson face_of_trans ff) + show ?thesis + using that + apply clarsimp + by (smt (verit, ccfv_SIG) IntI face_of_aff_dim_lt face_of_imp_convex [of X] face_of_imp_convex [of Y] face_of_trans ff) qed obtain g where "continuous_on (\(\ \ ?Faces)) g" "g ` \(\ \ ?Faces) \ rel_frontier T" @@ -711,10 +704,7 @@ have face': "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" by (metis face inf_commute) have *: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" - unfolding \_def - using subsetD [OF \\ \ \\] apply (auto simp add: face) - apply (meson face' face_of_Int_subface face_of_refl_eq poly polytope_imp_convex)+ - done + by (simp add: \_def) (smt (verit) \\ \ \\ DiffE face' face_of_Int_subface in_mono inf.idem) obtain h where conth: "continuous_on (\\) h" and him: "h ` (\\) \ rel_frontier T" and hf: "\x. x \ \\ \ h x = f x" proof (rule extend_map_lemma [OF \finite \\ [unfolded \_def] Un_upper1 T]) @@ -1143,7 +1133,7 @@ proof (cases "K = {}") case True then show ?thesis - by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that) + by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI image_eqI subset_iff that) next case False have "S \ U" @@ -1166,12 +1156,7 @@ by (simp_all add: in_components_subset comps that) then obtain a where a: "a \ C" "a \ L" by auto have opeUC: "openin (top_of_set U) C" - proof (rule openin_trans) - show "openin (top_of_set (U-S)) C" - by (simp add: \locally connected U\ clo locally_diff_closed openin_components_locally_connected [OF _ C]) - show "openin (top_of_set U) (U - S)" - by (simp add: clo openin_diff) - qed + by (metis C \locally connected U\ clo closedin_def locally_connected_open_component topspace_euclidean_subtopology) then obtain d where "C \ U" "0 < d" and d: "cball a d \ U \ C" using openin_contains_cball by (metis \a \ C\) then have "ball a d \ U \ C" @@ -1190,7 +1175,7 @@ show "finite (C \ K)" by (simp add: \finite K\) show "S \ C \ affine hull C" - by (metis \C \ U\ \S \ U\ \a \ C\ opeUC affine_hull_eq affine_hull_openin all_not_in_conv assms(2) sup.bounded_iff) + by (metis \S \ U\ \a \ C\ affine_hull_eq affine_hull_openin assms(2) empty_iff hull_subset le_sup_iff opeUC) show "connected C" by (metis C in_components_connected) qed auto @@ -1360,15 +1345,15 @@ proof (rule closedin_closed_subset [OF _ SU']) have *: "\C. C \ F \ openin (top_of_set U) C" unfolding F_def - by clarify (metis (no_types, lifting) \locally connected U\ clo closedin_def locally_diff_closed openin_components_locally_connected openin_trans topspace_euclidean_subtopology) + by (metis (no_types, lifting) \locally connected U\ clo closedin_def locally_connected_open_component mem_Collect_eq topspace_euclidean_subtopology) show "closedin (top_of_set U) (U - UF)" - unfolding UF_def - by (force intro: openin_delete *) - show "S \ \G = (U - UF) \ (S \ \G \ (S \ UF))" - using \S \ U\ apply (auto simp: F_def G_def UF_def) - apply (metis Diff_iff UnionI Union_components) - apply (metis DiffD1 UnionI Union_components) - by (metis (no_types, lifting) IntI components_nonoverlap empty_iff) + unfolding UF_def by (force intro: openin_delete *) + have "(\x\F. x - {a x}) \ S = {}" "\ G \ U" + using in_components_subset by (auto simp: F_def G_def) + moreover have "\ G \ UF = {}" + using components_nonoverlap by (fastforce simp: F_def G_def UF_def) + ultimately show "S \ \G = (U - UF) \ (S \ \G \ (S \ UF))" + using UF_def \S \ U\ by auto qed have clo2: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ UF)" proof (rule closedin_closed_subset [OF _ SU']) @@ -1378,13 +1363,14 @@ using F_def \locally connected U\ clo closedin_Un_complement_component by blast qed (simp add: \finite F\) show "S \ UF = (\C\F. S \ C) \ (S \ \G \ (S \ UF))" - using \S \ U\ apply (auto simp: F_def G_def UF_def) - using C0 apply blast - by (metis components_nonoverlap disjoint_iff) + proof + show "\ ((\) S ` F) \ (S \ \ G \ (S \ UF)) \ S \ UF" + using components_eq [of _ "U-S"] + by (auto simp add: F_def G_def UF_def disjoint_iff_not_equal) + qed (use UF_def \C0 \ F\ in blast) qed have SUG: "S \ \G \ U - K" - using \S \ U\ K apply (auto simp: G_def disjnt_iff) - by (meson Diff_iff subsetD in_components_subset) + using \S \ U\ K in_components_subset[of _ "U-S"] by (force simp: G_def disjnt_iff) then have contf': "continuous_on (S \ \G) f" by (rule continuous_on_subset [OF contf]) have contg': "continuous_on (S \ UF) g" @@ -1392,8 +1378,7 @@ have "\x. \S \ U; x \ S\ \ f x = g x" by (subst gh) (auto simp: ah C0 intro: \C0 \ F\) then have f_eq_g: "\x. x \ S \ UF \ x \ S \ \G \ f x = g x" - using \S \ U\ apply (auto simp: F_def G_def UF_def dest: in_components_subset) - using components_eq by blast + using \S \ U\ components_eq [of _ "U-S"] by (fastforce simp add: F_def G_def UF_def) have cont: "continuous_on (S \ \G \ (S \ UF)) (\x. if x \ S \ \G then f x else g x)" by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\x. x \ S \ \G"]) show ?thesis @@ -1553,8 +1538,7 @@ then have "\ bounded (\{C \ components (T - S). \ bounded C})" by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that) then show ?thesis - apply (clarsimp simp: LU_def Int_Un_distrib Diff_Int_distrib Int_UN_distrib) - by (metis (no_types, lifting) False Sup_upper bounded_cbox bounded_subset inf.orderE mem_Collect_eq that) + by (simp add: LU_def disjoint_iff) (meson False bounded_cbox bounded_subset subset_iff that) qed qed blast have *: False if "x \ cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)" @@ -1626,10 +1610,11 @@ then show ?thesis using cT_ne clo_cT closest_point_in_set by blast qed - show "continuous_on (T - K \ cbox (- (b + One)) (b + One)) (g \ closest_point (cbox (-c) c \ T))" - using cloTK - apply (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]) - by (auto simp add: clo_cT affine_imp_convex \affine T\ convex_Int cT_ne) + have "convex (cbox (- c) c \ T)" + by (simp add: affine_imp_convex assms(4) convex_Int) + then show "continuous_on (T - K \ cbox (- (b + One)) (b + One)) (g \ closest_point (cbox (-c) c \ T))" + using cloTK clo_cT cT_ne + by (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]; force) have "g (closest_point (cbox (- c) c \ T) x) \ rel_frontier U" if "x \ T" "x \ K \ x \ cbox (- b - One) (b + One)" for x using gim [THEN subsetD] that cloTK by blast @@ -1661,10 +1646,9 @@ then have "aff_dim T \ aff_dim (cball a r)" by (simp add: aff aff_dim_cball) then show ?thesis - apply (rule extend_map_affine_to_sphere_cofinite_gen - [OF \compact S\ convex_cball bounded_cball \affine T\ \S \ T\ _ contf]) - using fim apply (auto simp: assms False that dest: dis) - done + using extend_map_affine_to_sphere_cofinite_gen + [OF \compact S\ convex_cball bounded_cball \affine T\ \S \ T\ _ contf] + using fim by (fastforce simp: assms False that dest: dis) qed corollary extend_map_UNIV_to_sphere_cofinite: @@ -1688,9 +1672,9 @@ and fim: "f ` S \ sphere a r" and dis: "\C. C \ components(- S) \ \ bounded C" obtains g where "continuous_on UNIV g" "g ` UNIV \ sphere a r" "\x. x \ S \ g x = f x" - apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"]) - apply (auto dest: dis) -done + using extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"] + by (metis Compl_empty_eq dis subset_empty) + theorem Borsuk_separation_theorem_gen: fixes S :: "'a::euclidean_space set" @@ -1754,11 +1738,10 @@ qed next assume R: ?rhs - then show ?lhs - apply (simp add: Borsuk_separation_theorem_gen [OF \compact S\, symmetric]) - apply (auto simp: components_def connected_iff_eq_connected_component_set) - using connected_component_in apply fastforce - using cobounded_unique_unbounded_component [OF _ 2, of "-S"] \compact S\ compact_eq_bounded_closed by fastforce + then have "\c\components (- S). \ bounded c \ connected (- S)" + by (metis "2" assms(1) cobounded_has_bounded_component compact_imp_bounded double_complement) + with R show ?lhs + by (simp add: Borsuk_separation_theorem_gen [OF \compact S\]) qed @@ -2032,8 +2015,7 @@ show "open (k ` U)" by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope]) show "inj_on (k \ f \ h) (k ` U)" - apply (clarsimp simp: inj_on_def) - by (metis \U \ S\ fim homeomorphism_apply2 homhk image_subset_iff inj_onD injf subsetD) + by (smt (verit) comp_apply inj_on_def \U \ S\ fim homeomorphism_apply2 homhk image_iff injf subsetD) qed moreover have eq: "f ` U = h ` (k \ f \ h \ k) ` U" @@ -2117,8 +2099,7 @@ show "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) show "inj_on (h \ f) S" - apply (clarsimp simp: inj_on_def) - by (metis fim homeomorphism_apply2 [OF homkh] image_subset_iff inj_onD injf) + by (smt (verit, ccfv_SIG) comp_apply fim inj_on_def homeomorphism_apply2 [OF homkh] image_subset_iff injf) show "(h \ f) ` S \ U" using \V' \ U\ hfV' by auto qed (auto simp: assms) @@ -2149,8 +2130,7 @@ moreover have "(h \ f) ` S \ U" using \T \ U\ fim homeomorphism_image1 homhk by fastforce moreover have "inj_on (h \ f) S" - apply (clarsimp simp: inj_on_def) - by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf) + by (smt (verit, best) comp_apply inj_on_def fim homeomorphism_apply1 homhk image_subset_iff injf) ultimately have ope_hf: "openin (top_of_set U) ((h \ f) ` S)" using invariance_of_domain_subspaces [OF ope \subspace U\ \subspace U\] by blast have "(h \ f) ` S \ T" @@ -2313,9 +2293,13 @@ lemma injective_into_1d_imp_open_map_UNIV: fixes f :: "'a::euclidean_space \ real" assumes "open T" "continuous_on S f" "inj_on f S" "T \ S" - shows "open (f ` T)" - apply (rule invariance_of_domain_gen [OF \open T\]) - using assms by (auto simp: elim: continuous_on_subset subset_inj_on) + shows "open (f ` T)" +proof - + have "DIM(real) \ DIM('a)" + by simp + then show ?thesis + using invariance_of_domain_gen assms continuous_on_subset subset_inj_on by metis +qed lemma continuous_on_inverse_open: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" @@ -2417,7 +2401,7 @@ fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "open S" "S \ {}" "open T" "T \ {}" shows "DIM('a) = DIM('b)" - using assms + using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis inj_onI invariance_of_dimension) done @@ -2427,9 +2411,6 @@ assumes "S homeomorphic T" "interior S = {} \ interior T = {}" shows "(interior S) homeomorphic (interior T)" proof (cases "interior T = {}") - case True - with assms show ?thesis by auto -next case False then have "DIM('a) = DIM('b)" using assms @@ -2438,7 +2419,7 @@ done then show ?thesis by (rule homeomorphic_interiors_same_dimension [OF \S homeomorphic T\]) -qed +qed (use assms in auto) lemma homeomorphic_frontiers_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" @@ -2674,7 +2655,7 @@ obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" - using assms [unfolded homeomorphic_minimal] by auto + using assms by (auto simp: homeomorphic_minimal) have "aff_dim (affine hull S) \ aff_dim (affine hull T)" using False assms homeomorphic_aff_dim_le by blast moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" @@ -2777,16 +2758,14 @@ using Arg2pi_eq that h01 by (force simp: j_def) have eq: "S = S \ (UNIV \ {z. 0 \ Im z}) \ S \ (UNIV \ {z. Im z \ 0})" for S :: "(real*complex)set" by auto + have \
: "Arg2pi z \ 2 * pi" for z + by (simp add: Arg2pi order_le_less) have c1: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. 0 \ Im z}) (\x. h (fst x, Arg2pi (snd x) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) - apply (auto simp: Arg2pi) - apply (meson Arg2pi_lt_2pi linear not_le) - done + by (auto simp: Arg2pi \
) have c2: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. Im z \ 0}) (\x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) - apply (auto simp: Arg2pi) - apply (meson Arg2pi_lt_2pi linear not_le) - done + by (auto simp: Arg2pi \
) show "continuous_on ({0..1} \ sphere 0 1) j" apply (simp add: j_def) apply (subst eq) @@ -2815,10 +2794,11 @@ and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \ S" shows "homotopic_with_canon (\h. True) (sphere 0 1) S f g" proof - - have "homotopic_loops S (f \ exp \ (\t. of_real(2 * pi * t) * \)) (g \ exp \ (\t. of_real(2 * pi * t) * \))" - apply (rule S [unfolded simply_connected_homotopic_loops, rule_format]) - apply (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim) - done + let ?h = "(\t. complex_of_real (2 * pi * t) * \)" + have "homotopic_loops S (f \ exp \ ?h) (f \ exp \ ?h) \ homotopic_loops S (g \ exp \ ?h) (g \ exp \ ?h)" + by (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim) + then have "homotopic_loops S (f \ exp \ ?h) (g \ exp \ ?h)" + using S simply_connected_homotopic_loops by blast then show ?thesis apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps]) apply (auto simp: o_def complex_norm_eq_1_exp mult.commute) @@ -2833,9 +2813,7 @@ continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "\a. homotopic_with_canon (\h. True) (sphere 0 1) S h (\x. a)" - apply (rule_tac x="h 1" in exI) - apply (rule hom) - using assms by (auto) + by (metis conth continuous_on_const him hom image_subset_iff) lemma simply_connected_eq_homotopic_circlemaps2b: fixes S :: "'a::real_normed_vector set" @@ -2889,9 +2867,8 @@ continuous_on (sphere 0 1) f \ f ` (sphere 0 1) \ S \ continuous_on (sphere 0 1) g \ g ` (sphere 0 1) \ S \ homotopic_with_canon (\h. True) (sphere 0 1) S f g)" - apply (rule iffI) - apply (blast dest: simply_connected_eq_homotopic_circlemaps1) - by (simp add: simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3) + by (metis simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a + simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3) proposition simply_connected_eq_contractible_circlemap: fixes S :: "'a::real_normed_vector set" @@ -2900,9 +2877,8 @@ (\f::complex \ 'a. continuous_on (sphere 0 1) f \ f `(sphere 0 1) \ S \ (\a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)))" - apply (rule iffI) - apply (simp add: simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b) - using simply_connected_eq_homotopic_circlemaps3 by blast + by (metis simply_connected_eq_homotopic_circlemaps simply_connected_eq_homotopic_circlemaps2a + simply_connected_eq_homotopic_circlemaps3 simply_connected_imp_path_connected) corollary homotopy_eqv_simple_connectedness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" @@ -3019,8 +2995,12 @@ fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space" assumes "a \ 0" "c \ 0" shows "({x. a \ x = b} homeomorphic {x. c \ x = d} \ DIM('a) = DIM('b))" - apply (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms) - by (metis DIM_positive Suc_pred) +proof - + have "DIM('a) - Suc 0 = DIM('b) - Suc 0 \ DIM('a) = DIM('b)" + by (metis DIM_positive Suc_pred) + then show ?thesis + by (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms) +qed lemma homeomorphic_UNIV_UNIV: shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \ @@ -3043,9 +3023,7 @@ by (rule convex_cball) then obtain c where "homotopic_with_canon (\z. True) (sphere (0::complex) 1) (rel_frontier S) f (\x. c)" apply (rule inessential_spheremap_lowdim_gen [OF _ bounded_cball \convex S\ \bounded S\, where f=f]) - using f 3 - apply (auto simp: aff_dim_cball) - done + using f 3 by (auto simp: aff_dim_cball) then show "\a. homotopic_with_canon (\h. True) (sphere 0 1) (rel_frontier S) f (\x. a)" by blast qed @@ -3101,9 +3079,7 @@ by (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset) qed (use contf continuous_on_subset hgsub in blast) show "inj_on (f \ h) (g ` (S - {b}))" - using kjsub - apply (clarsimp simp add: inj_on_def) - by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD) + by (smt (verit, del_insts) SU homeomorphism_def inj_on_def injf gh Diff_iff comp_apply imageE subset_iff) show "(f \ h) ` g ` (S - {b}) \ T" by (metis fim image_comp image_mono hgsub subset_trans) qed (auto simp: assms) @@ -3118,9 +3094,7 @@ by (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset) qed (use contf continuous_on_subset kjsub in blast) show "inj_on (f \ k) (j ` (S - {c}))" - using kjsub - apply (clarsimp simp add: inj_on_def) - by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD) + by (smt (verit, del_insts) SU homeomorphism_def inj_on_def injf jk Diff_iff comp_apply imageE subset_iff) show "(f \ k) ` j ` (S - {c}) \ T" by (metis fim image_comp image_mono kjsub subset_trans) qed (auto simp: assms) @@ -3131,11 +3105,9 @@ have "h ` g ` (S - {b}) = (S - {b})" proof show "h ` g ` (S - {b}) \ S - {b}" - using homeomorphism_apply1 [OF gh] SU - by (fastforce simp add: image_iff image_subset_iff) + using homeomorphism_apply1 [OF gh] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {b} \ h ` g ` (S - {b})" - apply clarify - by (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def) + by (metis Diff_mono SU gh homeomorphism_image2 homeomorphism_of_subsets set_eq_subset) qed then show ?thesis by (metis image_comp) @@ -3143,14 +3115,8 @@ moreover have "(f \ k) ` j ` (S - {c}) = f ` (S - {c})" proof - have "k ` j ` (S - {c}) = (S - {c})" - proof - show "k ` j ` (S - {c}) \ S - {c}" - using homeomorphism_apply1 [OF jk] SU - by (fastforce simp add: image_iff image_subset_iff) - show "S - {c} \ k ` j ` (S - {c})" - apply clarify - by (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def) - qed + using homeomorphism_apply1 [OF jk] SU + by (meson Diff_mono homeomorphism_def homeomorphism_of_subsets jk subset_refl) then show ?thesis by (metis image_comp) qed @@ -3252,8 +3218,7 @@ using contractible_sphere contractible_def not_one_le_zero by blast qed with False show ?rhs - apply simp - by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3) + by (metis DIM_ge_Suc0 One_nat_def Suc_1 not_less_eq_eq numeral_3_eq_3 order_antisym_conv) next assume ?rhs with False show ?lhs by (simp add: simply_connected_sphere) @@ -3610,7 +3575,10 @@ using pi_ge_two by (simp add: ball_subset_ball_iff) then have inj_exp: "inj_on exp (ball (Ln z) 1)" using inj_on_exp_pi inj_on_subset by blast - define \ where "\ \ range (\n. (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1))" + define twopi where "twopi \ \n. of_real (2 * of_int n * pi) * \" + define \ where "\ \ range (\n. (\x. x + twopi n) ` (ball(Ln z) 1))" + have exp_eq': "exp w = exp z \ (\n::int. w = z + twopi n)" for z w + by (simp add: exp_eq twopi_def) show ?thesis proof (intro exI conjI) show "z \ exp ` (ball(Ln z) 1)" @@ -3619,80 +3587,73 @@ by blast with inj_exp show "openin (top_of_set (- {0})) (exp ` ball (Ln z) 1)" by (auto simp: openin_open_eq invariance_of_domain continuous_on_exp [OF continuous_on_id]) - show "\\ = exp -` exp ` ball (Ln z) 1" - by (force simp: \_def Complex_Transcendental.exp_eq image_iff) + show UV: "\\ = exp -` exp ` ball (Ln z) 1" + by (force simp: \_def twopi_def Complex_Transcendental.exp_eq image_iff) show "\V\\. open V" by (auto simp: \_def inj_on_def continuous_intros invariance_of_domain) - have xy: "2 \ cmod (2 * of_int x * of_real pi * \ - 2 * of_int y * of_real pi * \)" - if "x < y" for x y + have "2 \ cmod (twopi m -twopi n)" if "m \ n" for m n proof - - have "1 \ abs (x - y)" + have "1 \ abs (m - n)" using that by linarith - then have "1 \ cmod (of_int x - of_int y) * 1" + then have "1 \ cmod (of_int m - of_int n) * 1" by (metis mult.right_neutral norm_of_int of_int_1_le_iff of_int_abs of_int_diff) - also have "... \ cmod (of_int x - of_int y) * of_real pi" + also have "... \ cmod (of_int m - of_int n) * of_real pi" using pi_ge_two by (intro mult_left_mono) auto - also have "... \ cmod ((of_int x - of_int y) * of_real pi * \)" + also have "... \ cmod ((of_int m - of_int n) * of_real pi * \)" by (simp add: norm_mult) - also have "... \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" + also have "... \ cmod (of_int m * of_real pi * \ - of_int n * of_real pi * \)" by (simp add: algebra_simps) - finally have "1 \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" . - then have "2 * 1 \ cmod (2 * (of_int x * of_real pi * \ - of_int y * of_real pi * \))" + finally have "1 \ cmod (of_int m * of_real pi * \ - of_int n * of_real pi * \)" . + then have "2 * 1 \ cmod (2 * (of_int m * of_real pi * \ - of_int n * of_real pi * \))" by (metis mult_le_cancel_left_pos norm_mult_numeral1 zero_less_numeral) then show ?thesis - by (simp add: algebra_simps) + by (simp add: algebra_simps twopi_def) qed - show "disjoint \" - apply (clarsimp simp add: \_def pairwise_def disjnt_def add.commute [of _ "x*y" for x y] - ball_eq_ball_iff intro!: disjoint_ballI) - apply (auto simp: dist_norm neq_iff) - by (metis norm_minus_commute xy)+ + then show "disjoint \" + unfolding \_def pairwise_def disjnt_iff + by (smt (verit, best) add.commute add_diff_cancel_left' add_diff_eq dist_commute dist_complex_def dist_triangle imageE mem_ball) show "\u\\. \q. homeomorphism u (exp ` ball (Ln z) 1) exp q" proof fix u assume "u \ \" - then obtain n where n: "u = (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1)" + then obtain n where n: "u = (\x. x + twopi n) ` (ball(Ln z) 1)" by (auto simp: \_def) have "compact (cball (Ln z) 1)" by simp moreover have "continuous_on (cball (Ln z) 1) exp" by (rule continuous_on_exp [OF continuous_on_id]) moreover have "inj_on exp (cball (Ln z) 1)" - apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) - using pi_ge_two by (simp add: cball_subset_ball_iff) + using pi_ge_two inj_on_subset [OF inj_on_exp_pi [of "Ln z"]] + by (simp add: subset_iff) ultimately obtain \ where hom: "homeomorphism (cball (Ln z) 1) (exp ` cball (Ln z) 1) exp \" using homeomorphism_compact by blast have eq1: "exp ` u = exp ` ball (Ln z) 1" - apply (auto simp: algebra_simps n) - apply (rule_tac x = "_ + \ * (of_int n * (of_real pi * 2))" in image_eqI) - apply (auto simp: image_iff) - done - have \exp: "\ (exp x) + 2 * of_int n * of_real pi * \ = x" if "x \ u" for x + by (smt (verit) n exp_eq' image_cong image_image) + have \exp: "\ (exp x) + twopi n = x" if "x \ u" for x proof - - have "exp x = exp (x - 2 * of_int n * of_real pi * \)" - by (simp add: exp_eq) - then have "\ (exp x) = \ (exp (x - 2 * of_int n * of_real pi * \))" + have "exp x = exp (x - twopi n)" + using exp_eq' by auto + then have "\ (exp x) = \ (exp (x - twopi n))" by simp - also have "... = x - 2 * of_int n * of_real pi * \" + also have "... = x - twopi n" using \x \ u\ by (auto simp: n intro: homeomorphism_apply1 [OF hom]) finally show ?thesis by simp qed - have exp2n: "exp (\ (exp x) + 2 * of_int n * complex_of_real pi * \) = exp x" - if "dist (Ln z) x < 1" for x - using that by (auto simp: exp_eq homeomorphism_apply1 [OF hom]) + have exp2n: "exp (\ (exp x) + twopi n) = exp x" if "dist (Ln z) x < 1" for x + by (metis \exp exp_eq' imageI mem_ball n that) have "continuous_on (exp ` ball (Ln z) 1) \" by (meson ball_subset_cball continuous_on_subset hom homeomorphism_cont2 image_mono) - then have cont: "continuous_on (exp ` ball (Ln z) 1) (\x. \ x + 2 * of_int n * complex_of_real pi * \)" + then have cont: "continuous_on (exp ` ball (Ln z) 1) (\x. \ x + twopi n)" by (intro continuous_intros) - show "\q. homeomorphism u (exp ` ball (Ln z) 1) exp q" - apply (rule_tac x="(\x. x + of_real(2 * n * pi) * \) \ \" in exI) + have "homeomorphism u (exp ` ball (Ln z) 1) exp ((\x. x + twopi n) \ \)" unfolding homeomorphism_def apply (intro conjI ballI eq1 continuous_on_exp [OF continuous_on_id]) apply (auto simp: \exp exp2n cont n) apply (force simp: image_iff homeomorphism_apply1 [OF hom])+ done + then show "\q. homeomorphism u (exp ` ball (Ln z) 1) exp q" by metis qed qed qed @@ -3787,7 +3748,7 @@ using hom by (auto simp: homotopic_with_def) show ?thesis apply (simp add: homotopic_with) - apply (rule_tac x="\z. k z*(h \ snd)z" in exI) + apply (rule_tac x="\z. k z * (h \ snd)z" in exI) using kim hin by (fastforce simp: conth norm_mult k0 k1 intro!: contk continuous_intros)+ qed @@ -3832,10 +3793,10 @@ have [simp]: "\x. x \ S \ g x \ 0" using geq1 by fastforce have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" - apply (rule homotopic_with_eq [OF homotopic_with_sphere_times [OF L cont]]) + using homotopic_with_eq [OF homotopic_with_sphere_times [OF L cont]] by (auto simp: divide_inverse norm_inverse) with L show ?rhs - by (auto simp: homotopic_with_imp_continuous dest: homotopic_with_imp_subset1 homotopic_with_imp_subset2) + by (simp add: homotopic_with_imp_continuous homotopic_with_imp_subset1 homotopic_with_imp_subset2) next assume ?rhs then show ?lhs by (elim conjE homotopic_with_eq [OF homotopic_with_sphere_times]; force) @@ -3998,15 +3959,13 @@ by (force simp: openin_euclidean_subtopology_iff dest: fST) have oo: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}" - apply (rule lower_hemicontinuous [THEN iffD1, rule_format]) - using fST clo by auto + using lower_hemicontinuous fST clo by blast have "compact (closure(f x))" by (simp add: bofx) moreover have "closure(f x) \ (\a \ f x. ball a (e/2))" using \0 < e\ by (force simp: closure_approachable simp del: divide_const_simps) ultimately obtain C where "C \ f x" "finite C" "closure(f x) \ (\a \ C. ball a (e/2))" - apply (rule compactE, force) - by (metis finite_subset_image) + by (meson compactE finite_subset_image Elementary_Metric_Spaces.open_ball compactE_image) then have fx_cover: "f x \ (\a \ C. ball a (e/2))" by (meson closure_subset order_trans) with fx_ne have "C \ {}" @@ -4173,9 +4132,10 @@ qed auto then have "homotopic_with_canon (\f. True) (sphere a r) (sphere b s) (k \ (h \ f)) (k \ (\x. c))" by (rule homotopic_with_compose_continuous_left [OF _ contk kim]) - then have "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. k c)" - apply (rule homotopic_with_eq, auto) + moreover have "\x. r = dist a x \ f x = k (h (f x))" by (metis fim hk homeomorphism_def image_subset_iff mem_sphere) + ultimately have "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. k c)" + by (auto intro: homotopic_with_eq) then show ?thesis by (metis that) qed @@ -4318,9 +4278,8 @@ by (simp add: exp_minus fgeq that) finally have "f z = exp (- h z) + \*g z" by (simp add: diff_eq_eq) - then show ?thesis - apply (simp add: cos_exp_eq) - by (metis fgeq add.assoc mult_2_right that) + with that show ?thesis + by (simp add: cos_exp_eq flip: fgeq) qed qed qed @@ -4391,11 +4350,7 @@ qed lemma retract_of_Borsukian: "\Borsukian T; S retract_of T\ \ Borsukian S" - apply (auto simp: retract_of_def retraction_def) - apply (erule (1) Borsukian_retraction_gen) - apply (meson retraction retraction_def) - apply (auto) - done + by (smt (verit) Borsukian_retraction_gen retract_of_def retraction retraction_def retraction_subset) lemma homeomorphic_Borsukian: "\Borsukian S; S homeomorphic T\ \ Borsukian T" using Borsukian_retraction_gen order_refl @@ -4580,12 +4535,7 @@ have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" - proof (intro continuous_intros) - show "continuous_on (S \ T) g" - by (meson contg continuous_on_subset inf_le1) - show "continuous_on (S \ T) h" - by (meson conth continuous_on_subset inf_sup_ord(2)) - qed + by (metis contg conth continuous_on_diff continuous_on_subset inf_le1 inf_le2) show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - @@ -4763,7 +4713,7 @@ by (simp add: Abstract_Topology_2.continuous_imp_closed_map \compact S\ contf fim) show "\U. closedin (top_of_set S) U \ closedin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" - using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]] + using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]] by meson show "bounded {z \ S. f z = y}" by (metis (no_types, lifting) compact_imp_bounded [OF \compact S\] bounded_subset mem_Collect_eq subsetI) @@ -4783,7 +4733,7 @@ by (fastforce simp: dist_norm abs_diff_less_iff algebra_simps) qed then show "\d>0. \x'\T. dist x' y < d \ dist (h (k x')) (h (k y)) < e" - using \0 < \\ by (auto simp: dist_commute) + using \0 < \\ by (auto simp: dist_commute) qed then show "\h. continuous_on T h \ (\x\T. g x = exp (\ * complex_of_real (h x)))" using fk gfh kTS by force @@ -5537,10 +5487,6 @@ by (simp_all add: assms) show "homotopic_with_canon (\x. True) S T f (\x. a)" if "S \ \" for S proof (cases "S = {}") - case True - then show ?thesis - by auto - next case False then obtain b where "b \ S" by blast @@ -5549,11 +5495,10 @@ then have "c \ T" using \b \ S\ homotopic_with_imp_subset2 by blast then have "homotopic_with_canon (\x. True) S T (\x. a) (\x. c)" - using T \a \ T\ homotopic_constant_maps path_connected_component - by (simp add: homotopic_constant_maps path_connected_component) + using T \a \ T\ by (simp add: homotopic_constant_maps path_connected_component) then show ?thesis using c homotopic_with_symD homotopic_with_trans by blast - qed + qed auto qed then show ?thesis .. qed @@ -5568,7 +5513,7 @@ by (simp add: assms compact_Un) with Borsukian_imp_unicoherent [of "S \ T"] ST assms show ?thesis - by (auto simp: closed_subset compact_imp_closed Borsukian_separation_compact unicoherent_def) + by (simp add: Borsukian_separation_compact closed_subset compact_imp_closed unicoherentD) qed end diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/BNF_Cardinal_Arithmetic.thy --- a/src/HOL/BNF_Cardinal_Arithmetic.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/BNF_Cardinal_Arithmetic.thy Sun Jan 15 20:00:44 2023 +0100 @@ -8,35 +8,37 @@ section \Cardinal Arithmetic as Needed by Bounded Natural Functors\ theory BNF_Cardinal_Arithmetic -imports BNF_Cardinal_Order_Relation + imports BNF_Cardinal_Order_Relation begin 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) + by (rule dir_image_ordIso) (auto simp add: inj_on_def card_order_on_def) 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 + have "Field (dir_image r f) = UNIV" + using assms 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 + using card_order_on_Card_order Card_order_ordIso2[OF _ dir_image] by blast ultimately show ?thesis by auto qed lemma ordIso_refl: "Card_order r \ r =o r" -by (rule card_order_on_ordIso) + by (rule card_order_on_ordIso) lemma ordLeq_refl: "Card_order r \ r \o r" -by (rule ordIso_imp_ordLeq, rule card_order_on_ordIso) + by (rule ordIso_imp_ordLeq, rule card_order_on_ordIso) lemma card_of_ordIso_subst: "A = B \ |A| =o |B|" -by (simp only: ordIso_refl card_of_Card_order) + by (simp only: ordIso_refl card_of_Card_order) lemma Field_card_order: "card_order r \ Field r = UNIV" -using card_order_on_Card_order[of UNIV r] by simp + using card_order_on_Card_order[of UNIV r] by simp subsection \Zero\ @@ -44,13 +46,12 @@ definition czero where "czero = card_of {}" -lemma czero_ordIso: - "czero =o czero" -using card_of_empty_ordIso by (simp add: czero_def) +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) + 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 @@ -62,28 +63,21 @@ lemma czeroI: "\Card_order r; Field r = {}\ \ r =o czero" -using Cnotzero_imp_not_empty ordIso_transitive[OF _ czero_ordIso] by blast + 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) + 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 + by (force intro: czeroI dest: card_of_mono2 card_of_empty3 czeroE) subsection \(In)finite cardinals\ definition cinfinite where - "cinfinite r = (\ finite (Field r))" + "cinfinite r \ (\ finite (Field r))" abbreviation Cinfinite where "Cinfinite r \ cinfinite r \ Card_order r" @@ -101,7 +95,7 @@ lemmas natLeq_card_order = natLeq_Card_order[unfolded Field_natLeq] lemma natLeq_cinfinite: "cinfinite natLeq" -unfolding cinfinite_def Field_natLeq by (rule infinite_UNIV_nat) + unfolding cinfinite_def Field_natLeq by (rule infinite_UNIV_nat) lemma natLeq_Cinfinite: "Cinfinite natLeq" using natLeq_cinfinite natLeq_Card_order by simp @@ -117,50 +111,46 @@ qed lemma cinfinite_not_czero: "cinfinite r \ \ (r =o (czero :: 'a rel))" -unfolding cinfinite_def by (cases "Field r = {}") (auto dest: czeroE) + unfolding cinfinite_def by (cases "Field r = {}") (auto dest: czeroE) lemma Cinfinite_Cnotzero: "Cinfinite r \ Cnotzero r" -by (rule conjI[OF cinfinite_not_czero]) simp_all + using cinfinite_not_czero by auto lemma Cinfinite_cong: "\r1 =o r2; Cinfinite r1\ \ Cinfinite r2" -using Card_order_ordIso2[of r1 r2] unfolding cinfinite_def ordIso_iff_ordLeq -by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2]) + using Card_order_ordIso2[of r1 r2] unfolding cinfinite_def ordIso_iff_ordLeq + by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2]) lemma cinfinite_mono: "\r1 \o r2; cinfinite r1\ \ cinfinite r2" -unfolding cinfinite_def by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2]) + unfolding cinfinite_def by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2]) lemma regularCard_ordIso: -assumes "k =o k'" and "Cinfinite k" and "regularCard k" -shows "regularCard k'" + assumes "k =o k'" and "Cinfinite k" and "regularCard k" + shows "regularCard k'" proof- have "stable k" using assms cinfinite_def regularCard_stable by blast hence "stable k'" using assms stable_ordIso1 ordIso_symmetric by blast - thus ?thesis using assms cinfinite_def stable_regularCard - using Cinfinite_cong by blast + thus ?thesis using assms cinfinite_def stable_regularCard Cinfinite_cong by blast qed corollary card_of_UNION_ordLess_infinite_Field_regularCard: -assumes ST: "regularCard r" and INF: "Cinfinite r" and - LEQ_I: "|I| i \ I. |A i| i \ I. A i| i \ I. |A i| i \ I. A i| Binary sum\ -definition csum (infixr "+c" 65) where - "r1 +c r2 \ |Field r1 <+> Field r2|" +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 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 using Cnotzero_imp_not_empty[of r1] Plus_eq_empty_conv[of "Field r1" "Field r2"] - card_of_ordIso_czero_iff_empty[of "Field r1 <+> Field r2"] by (auto intro: card_of_Card_order) +lemma csum_Cnotzero1: "Cnotzero r1 \ Cnotzero (r1 +c r2)" + using Cnotzero_imp_not_empty + by (auto intro: card_of_Card_order simp: csum_def card_of_ordIso_czero_iff_empty) lemma card_order_csum: assumes "card_order r1" "card_order r2" @@ -172,49 +162,49 @@ 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 (rule conjI[OF _ card_of_Card_order]) (auto simp: Field_card_of) + unfolding cinfinite_def csum_def by (auto simp: Field_card_of) lemma Cinfinite_csum: "Cinfinite r1 \ Cinfinite r2 \ Cinfinite (r1 +c r2)" -unfolding cinfinite_def csum_def by (rule conjI[OF _ card_of_Card_order]) (auto simp: Field_card_of) + using card_of_Card_order + by (auto simp: cinfinite_def csum_def Field_card_of) + +lemma Cinfinite_csum1: "Cinfinite r1 \ Cinfinite (r1 +c r2)" + by (blast intro!: Cinfinite_csum elim: ) lemma Cinfinite_csum_weak: "\Cinfinite r1; Cinfinite r2\ \ Cinfinite (r1 +c r2)" by (erule Cinfinite_csum1) lemma csum_cong: "\p1 =o r1; p2 =o r2\ \ p1 +c p2 =o r1 +c r2" -by (simp only: csum_def ordIso_Plus_cong) + 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) + 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) + 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) + 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) + 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) + 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) + 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) + 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) + 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) + 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 @@ -235,10 +225,10 @@ qed lemma Plus_csum: "|A <+> B| =o |A| +c |B|" -by (simp only: csum_def Field_card_of card_of_refl) + 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 + using ordLeq_ordIso_trans[OF card_of_Un_Plus_ordLeq Plus_csum] by blast subsection \One\ @@ -246,16 +236,17 @@ "cone = card_of {()}" lemma Card_order_cone: "Card_order cone" -unfolding cone_def by (rule card_of_Card_order) + 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 ordIso_iff_ordLeq using card_of_empty3 empty_not_insert by blast + unfolding czero_def cone_def ordIso_iff_ordLeq + using card_of_empty3 empty_not_insert by blast lemma cone_ordLeq_Cnotzero: "Cnotzero r \ cone \o r" -unfolding cone_def by (rule Card_order_singl_ordLeq) (auto intro: czeroI) + unfolding cone_def by (rule Card_order_singl_ordLeq) (auto intro: czeroI) subsection \Two\ @@ -264,14 +255,14 @@ "ctwo = |UNIV :: bool set|" lemma Card_order_ctwo: "Card_order ctwo" -unfolding ctwo_def by (rule card_of_Card_order) + 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 using UNIV_not_empty by auto + using card_of_empty3[of "UNIV :: bool set"] ordIso_iff_ordLeq + unfolding czero_def ctwo_def using UNIV_not_empty by auto lemma ctwo_Cnotzero: "Cnotzero ctwo" -by (simp add: ctwo_not_czero Card_order_ctwo) + by (simp add: ctwo_not_czero Card_order_ctwo) subsection \Family sum\ @@ -288,7 +279,7 @@ "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) + by (auto simp: Csum_def Field_card_of) (* NB: Always, under the cardinal operator, operations on sets are reduced automatically to operations on cardinals. @@ -304,49 +295,50 @@ 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 + 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) + 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) + 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) + by (simp only: cprod_def ordLeq_Times_mono2) lemma cprod_mono: "\p1 \o r1; p2 \o r2\ \ p1 *c p2 \o r1 *c r2" -by (rule ordLeq_transitive[OF cprod_mono1 cprod_mono2]) + by (rule ordLeq_transitive[OF cprod_mono1 cprod_mono2]) lemma ordLeq_cprod2: "\Cnotzero p1; Card_order p2\ \ p2 \o p1 *c p2" -unfolding cprod_def by (rule Card_order_Times2) (auto intro: czeroI) + unfolding cprod_def by (rule Card_order_Times2) (auto intro: 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) + 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 (rule cinfinite_mono) (auto intro: ordLeq_cprod2) + by (rule cinfinite_mono) (auto intro: ordLeq_cprod2) lemma Cinfinite_cprod2: "\Cnotzero r1; Cinfinite r2\ \ Cinfinite (r1 *c r2)" -by (blast intro: cinfinite_cprod2 Card_order_cprod) + by (blast intro: cinfinite_cprod2 Card_order_cprod) lemma cprod_cong: "\p1 =o r1; p2 =o r2\ \ p1 *c p2 =o r1 *c r2" -unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono) + unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono) lemma cprod_cong1: "\p1 =o r1\ \ p1 *c p2 =o r1 *c p2" -unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono1) + unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono1) lemma cprod_cong2: "p2 =o r2 \ q *c p2 =o q *c r2" -unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono2) + unfolding ordIso_iff_ordLeq by (blast intro: cprod_mono2) lemma cprod_com: "p1 *c p2 =o p2 *c p1" -by (simp only: cprod_def card_of_Times_commute) + 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_mono1) + by (simp only: Csum_def cprod_def Field_card_of card_of_Sigma_mono1) lemma card_of_Csum_Times': assumes "Card_order r" "\i \ I. |A i| \o r" @@ -361,27 +353,33 @@ 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) + 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 (rule conjunct2[OF Card_order_Plus_infinite]) - (auto simp: cinfinite_def dest: cinfinite_mono) + unfolding csum_def + using Card_order_Plus_infinite + by (fastforce simp: cinfinite_def dest: cinfinite_mono) lemma csum_absorb1': assumes card: "Card_order r2" - and r12: "r1 \o r2" and cr12: "cinfinite r1 \ cinfinite 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)+) +proof - + have "r1 +c r2 =o r2" + by (simp add: csum_absorb2' assms) + then show ?thesis + by (blast intro: ordIso_transitive csum_com) +qed lemma csum_absorb1: "\Cinfinite r2; r1 \o r2\ \ r2 +c r1 =o r2" -by (rule csum_absorb1') auto + by (rule csum_absorb1') auto lemma csum_absorb2: "\Cinfinite r2 ; r1 \o r2\ \ r1 +c r2 =o r2" using ordIso_transitive csum_com csum_absorb1 by blast lemma regularCard_csum: assumes "Cinfinite r" "Cinfinite s" "regularCard r" "regularCard s" - shows "regularCard (r +c s)" + shows "regularCard (r +c s)" proof (cases "r \o s") case True then show ?thesis using regularCard_ordIso[of s] csum_absorb2'[THEN ordIso_symmetric] assms by auto @@ -394,9 +392,9 @@ lemma csum_mono_strict: assumes Card_order: "Card_order r" "Card_order q" - and Cinfinite: "Cinfinite r'" "Cinfinite q'" - and less: "r o q" using finite_ordLess_infinite cinfinite_def cfinite_def outer Well_order - ordLess_imp_ordLeq by blast + ordLess_imp_ordLeq by blast then have "r +c q =o q" by (rule csum_absorb2[OF True]) then show ?thesis using ordIso_ordLess_trans ordLess_ordLeq_trans less ordLeq_csum2 Cinfinite by blast next @@ -453,11 +451,11 @@ "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) + 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 = {}" + and n: "Field p2 = {} \ Field r2 = {}" shows "p1 ^c p2 \o r1 ^c r2" proof(cases "Field p1 = {}") case True @@ -498,36 +496,36 @@ 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" + and n: "p2 =o czero \ r2 =o czero" and card: "Card_order p2" shows "p1 ^c p2 \o r1 ^c r2" by (rule cexp_mono'[OF 1 2 czeroE[OF n[OF czeroI[OF card]]]]) 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) + 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 = {}" + 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 + 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" + 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 + 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" -using assms(3) czeroI by (blast intro: cexp_mono2'[OF assms(1,2)]) + using assms(3) czeroI by (blast intro: cexp_mono2'[OF assms(1,2)]) lemma cexp_cong: assumes 1: "p1 =o r1" and 2: "p2 =o r2" - and Cr: "Card_order r2" - and Cp: "Card_order p2" + 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)" @@ -535,7 +533,7 @@ 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 + 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 @@ -543,12 +541,12 @@ 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]) + 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) + 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" @@ -570,31 +568,30 @@ 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) + using cprod_com r1 by (intro cexp_cong2, 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)+ + unfolding cinfinite_def cprod_def + by (rule Card_order_Times_infinite[THEN conjunct1]) (blast intro: czeroI)+ lemma cprod_infinite: "Cinfinite r \ r *c r =o r" -using cprod_infinite1' Cinfinite_Cnotzero ordLeq_refl by blast + using cprod_infinite1' Cinfinite_Cnotzero ordLeq_refl by blast lemma cexp_cprod_ordLeq: assumes r1: "Card_order r1" and r2: "Cinfinite r2" - and r3: "Cnotzero r3" "r3 \o 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)+ + using assms by (fastforce simp: Card_order_cprod intro: cprod_infinite1' cexp_cong2) finally show ?thesis . qed lemma Cnotzero_UNIV: "Cnotzero |UNIV|" -by (auto simp: card_of_Card_order card_of_ordIso_czero_iff_empty) + by (auto simp: card_of_Card_order card_of_ordIso_czero_iff_empty) lemma ordLess_ctwo_cexp: assumes "Card_order r" @@ -613,21 +610,12 @@ 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 + have "q =o q ^c cone" + by (blast intro: assms ordIso_symmetric cexp_cone) + also have "q ^c cone \o q ^c r" + using assms + by (intro cexp_mono2) (simp_all add: cone_ordLeq_Cnotzero cone_not_czero Card_order_cone) + finally show ?thesis . qed lemma ordLeq_cexp2: @@ -636,24 +624,20 @@ 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 + case False + have "r o q ^c r" + by (blast intro: assms cexp_mono1) + finally show ?thesis by (rule ordLess_imp_ordLeq) qed lemma cinfinite_cexp: "\ctwo \o q; Cinfinite r\ \ cinfinite (q ^c r)" -by (rule cinfinite_mono[OF ordLeq_cexp2]) simp_all + by (rule cinfinite_mono[OF ordLeq_cexp2]) simp_all lemma Cinfinite_cexp: "\ctwo \o q; Cinfinite r\ \ Cinfinite (q ^c r)" -by (simp add: cinfinite_cexp Card_order_cexp) + by (simp add: cinfinite_cexp Card_order_cexp) lemma card_order_cexp: assumes "card_order r1" "card_order r2" @@ -664,32 +648,32 @@ qed lemma ctwo_ordLess_natLeq: "ctwo ctwo o r" -by (rule ordLess_imp_ordLeq[OF ctwo_ordLess_Cinfinite[OF assms]]) + 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) + by (auto simp add: cinfinite_def card_of_Un_ordLeq_infinite_Field) lemma Un_Cinfinite_bound_strict: "\|A| \ |A \ B| |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) + 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+ + have "|Field p| \o r" "|Field q| \o r" + using assms card_of_least ordLeq_transitive unfolding card_order_on_def by blast+ with assms show ?thesis unfolding cinfinite_def csum_def by (blast intro: card_of_Plus_ordLeq_infinite_Field) qed @@ -711,26 +695,22 @@ lemma regularCard_cprod: assumes "Cinfinite r" "Cinfinite s" "regularCard r" "regularCard s" - shows "regularCard (r *c s)" + shows "regularCard (r *c s)" proof (cases "r \o s") case True - show ?thesis - apply (rule regularCard_ordIso[of s]) - apply (rule ordIso_symmetric[OF cprod_infinite2']) - using assms True Cinfinite_Cnotzero by auto + with assms Cinfinite_Cnotzero show ?thesis + by (force intro: regularCard_ordIso ordIso_symmetric[OF cprod_infinite2']) next case False have "Well_order r" "Well_order s" using assms card_order_on_well_order_on by auto - then have 1: "s \o r" using not_ordLeq_iff_ordLess ordLess_imp_ordLeq False by blast - show ?thesis - apply (rule regularCard_ordIso[of r]) - apply (rule ordIso_symmetric[OF cprod_infinite1']) - using assms 1 Cinfinite_Cnotzero by auto + then have "s \o r" using not_ordLeq_iff_ordLess ordLess_imp_ordLeq False by blast + with assms Cinfinite_Cnotzero show ?thesis + by (force intro: regularCard_ordIso ordIso_symmetric[OF cprod_infinite1']) 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 + 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") @@ -742,8 +722,8 @@ 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]) + 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) @@ -807,12 +787,12 @@ (* cardSuc *) lemma Cinfinite_cardSuc: "Cinfinite r \ Cinfinite (cardSuc r)" -by (simp add: cinfinite_def cardSuc_Card_order cardSuc_finite) + by (simp add: cinfinite_def cardSuc_Card_order cardSuc_finite) lemma cardSuc_UNION_Cinfinite: assumes "Cinfinite r" "relChain (cardSuc r) As" "B \ (\i \ Field (cardSuc r). As i)" "|B| <=o r" shows "\i \ Field (cardSuc r). B \ As i" -using cardSuc_UNION assms unfolding cinfinite_def by blast + using cardSuc_UNION assms unfolding cinfinite_def by blast lemma Cinfinite_card_suc: "\ Cinfinite r ; card_order r \ \ Cinfinite (card_suc r)" using Cinfinite_cong[OF cardSuc_ordIso_card_suc Cinfinite_cardSuc] . diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/BNF_Cardinal_Order_Relation.thy --- a/src/HOL/BNF_Cardinal_Order_Relation.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/BNF_Cardinal_Order_Relation.thy Sun Jan 15 20:00:44 2023 +0100 @@ -9,7 +9,7 @@ section \Cardinal-Order Relations as Needed by Bounded Natural Functors\ theory BNF_Cardinal_Order_Relation -imports Zorn BNF_Wellorder_Constructions + imports Zorn BNF_Wellorder_Constructions begin text\In this section, we define cardinal-order relations to be minim well-orders @@ -42,20 +42,20 @@ strict order-embedding relation, \), among all the well-orders on its field.\ definition card_order_on :: "'a set \ 'a rel \ bool" -where -"card_order_on A r \ well_order_on A r \ (\r'. well_order_on A r' \ r \o r')" + 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 + 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 well_order_on_Field by blast + "card_order_on A r \ A = Field r \ Card_order r" + unfolding card_order_on_def using 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: @@ -68,48 +68,47 @@ \ theorem card_order_on: "\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 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 +proof - + define R where "R \ {r. well_order_on A r}" + have "R \ {} \ (\r \ R. Well_order r)" + using well_order_on[of A] R_def well_order_on_Well_order by blast + with exists_minim_Well_order[of R] show ?thesis + by (auto simp: R_def card_order_on_def) 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 + 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 + 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 well_order_on_Well_order by blast + using 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 + 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) + 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_Field) + 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_Field) 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 + 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 + 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\ @@ -119,220 +118,213 @@ 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)" + 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) + 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 + 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 well_order_on_Field by blast + using card_of_card_order_on[of A] unfolding card_order_on_def + using 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) + 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 + "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 + 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 + 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 + 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| )" + "(\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 + 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 + 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 + 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 + 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| )" + "(\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 Schroeder_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 + hence "|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 Schroeder_Bernstein[of f] by fastforce + hence "|A| \o |B|" using card_of_ordIso ordIso_iff_ordLeq by auto } thus "|A| \o |B|" using ordLess_or_ordLeq[of "|B|" "|A|"] - by (auto simp: card_of_Well_order) + 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 + 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|"] card_of_Well_order embed_Field[of "|A|" "|B|"] + by (auto simp: Field_card_of) 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 + "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)) = ( |B| (\f. inj_on f A \ f ` A \ B)) = (\ |A| \o |B| )" - using card_of_ordLeq by blast + using card_of_ordLeq by blast also have "\ = ( |B| {} \ (\(\f. f ` A = B)) = ( |A| {} \ (\(\f. f ` A = B)) = ( |A| a. a \ A \ f a \ B" -shows "|A| \o |B|" -using assms unfolding card_of_ordLeq[symmetric] by auto + assumes "inj_on f A" and "\ 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) + "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 + "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- + 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: well_order_on_Well_order) + 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: 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) + 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 well_order_on_Well_order by blast + by (simp add: ordIso_iff_ordLeq card_of_mono2) lemma card_of_Field_ordIso: -assumes "Card_order r" -shows "|Field r| =o r" -proof- + 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 + 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 + 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 + "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- + "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) + 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_least + by (auto simp: card_of_least 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 (underS r a) a \ Field r. Restr r (underS r a) Field r" -shows "|underS r a| Field r" + shows "|underS r a| o ?r'" - unfolding card_order_on_def by simp + with card_of_card_order_on have "|Field ?r'| \o ?r'" + unfolding card_order_on_def by auto moreover have "Field ?r' = ?A" - using 1 wo_rel.underS_ofilter Field_Restr_ofilter - unfolding wo_rel_def by fastforce + 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)" + "( |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 + 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 + 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" @@ -340,8 +332,8 @@ 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 + "( |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\ @@ -353,123 +345,113 @@ \ lemma card_of_empty: "|{}| \o |A|" -using card_of_ordLeq inj_on_id by blast + using card_of_ordLeq inj_on_id by blast lemma card_of_empty1: -assumes "Well_order r \ Card_order r" -shows "|{}| \o r" -proof- + 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 + hence "|Field r| \o r" + using assms card_of_least 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) + "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 + assumes "|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) + assumes "|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 + "|{}::'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 \ {}" + "|f ` A| \o |A|" +proof(cases "A = {}") + case False hence "f ` A \ {}" by auto - thus "|f ` A| \o |A|" - using card_of_ordLeq2[of "f ` A" A] by auto -qed + thus ?thesis + using card_of_ordLeq2[of "f ` A" A] by auto +qed (simp add: card_of_empty) 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 + 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_singl_ordLeq: -assumes "A \ {}" -shows "|{b}| \o |A|" -proof- + 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 + using * unfolding inj_on_def by auto thus ?thesis unfolding card_of_ordLeq[symmetric] by (intro exI) 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 + "\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 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 +lemma card_of_Plus1: "|A| \o |A <+> B|" and card_of_Plus2: "|B| \o |A <+> B|" + using card_of_ordLeq by force+ 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 + "Card_order r \ r \o |(Field r) <+> B|" + using card_of_Plus1 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast 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 + "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- +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- +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" +proof - + let ?f = "\c. 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 + 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|" + fixes A :: "'a set" and B :: "'b set" and C :: "'c set" + shows "|(A <+> B) <+> C| =o |A <+> B <+> C|" proof - define f :: "('a + 'b) + 'c \ 'a + 'b + 'c" where [abs_def]: "f k = @@ -487,21 +469,11 @@ proof(cases x) case (Inl a) hence "a \ A" "x = f (Inl (Inl a))" - using x unfolding f_def by auto + 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 + case (Inr bc) with x show ?thesis + by (cases bc) (force simp: f_def)+ qed qed hence "bij_betw f ((A <+> B) <+> C) (A <+> B <+> C)" @@ -510,159 +482,134 @@ 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 + assumes "|A| \o |B|" + shows "|A <+> C| \o |B <+> C|" +proof - + obtain f where f: "inj_on f A \ f ` A \ B" + using assms card_of_ordLeq[of A] by fastforce + define g where "g \ \d. case d of Inl a \ Inl(f a) | Inr (c::'c) \ Inr c" 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 g_def by force - } - moreover - {fix d assume "d \ A <+> C" - hence "g d \ B <+> C" using 1 - by(cases d) (auto simp add: g_def) - } - ultimately show ?thesis unfolding inj_on_def by auto - qed + using f unfolding inj_on_def g_def by force thus ?thesis using card_of_ordLeq by blast 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 + 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 + assumes "|A| \o |B|" + shows "|C <+> A| \o |C <+> B|" + using card_of_Plus_mono1[OF assms] + by (blast intro: card_of_Plus_commute ordIso_ordLeq_trans ordLeq_ordIso_trans) 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 + 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 + 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 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 + 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) + 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 + 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) + 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 + 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) + 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 + 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 + "|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 \ {}" + assumes "A \ {}" + shows "|B| \o |B \ A|" +proof(cases "B = {}") + case False have "fst `(B \ A) = B" 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 + card_of_ordLeq False by blast +qed (simp add: card_of_empty) 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 +proof - + have "bij_betw (\(a,b). (b,a)) (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 + 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 + "\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 + "\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) + using card_of_Times1[of A] + by(cases "A = {}", simp add: card_of_empty) lemma card_of_Plus_Times_bool: "|A <+> A| =o |A \ (UNIV::bool set)|" -proof- +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(cases c1; cases c2) auto - } + proof - + have "\c1 c2. ?f c1 = ?f c2 \ c1 = c2" + by (force split: sum.split_asm) moreover - {fix c assume "c \ A <+> A" - hence "?f c \ A \ (UNIV::bool set)" - by(cases c) auto - } + have "\c. c \ A <+> A \ ?f c \ A \ (UNIV::bool set)" + by (force split: sum.split_asm) 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 + {fix a bl assume "(a,bl) \ A \ (UNIV::bool set)" + hence "(a,bl) \ ?f ` ( A <+> A)" + by (cases bl) (force split: sum.split_asm)+ } ultimately show ?thesis unfolding bij_betw_def inj_on_def by auto qed @@ -670,127 +617,107 @@ 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 + assumes "|A| \o |B|" + shows "|A \ C| \o |B \ C|" +proof - + obtain f where f: "inj_on f A \ f ` A \ B" + using assms card_of_ordLeq[of A] by fastforce + define g where "g \ (\(a,c::'c). (f a,c))" have "inj_on g (A \ C) \ g ` (A \ C) \ (B \ C)" - using 1 unfolding inj_on_def using g_def by auto + using f unfolding inj_on_def using g_def by auto thus ?thesis using card_of_ordLeq by blast 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 + 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 + assumes "|A| \o |B|" + shows "|C \ A| \o |C \ B|" + using assms card_of_Times_mono1[of A B C] + by (blast intro: card_of_Times_commute ordIso_ordLeq_trans ordLeq_ordIso_trans) 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 + 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- + 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) + 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" + obtain F where F: "\i \ I. inj_on (F i) (A i) \ (F i) ` (A i) \ B i" by atomize_elim (auto intro: bchoice) - obtain g where g_def: "g = (\(i,a::'b). (i,F i a))" by blast + define g where "g \ (\(i,a::'b). (i,F i a))" 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 + using F unfolding inj_on_def using g_def by force thus ?thesis using card_of_ordLeq by blast qed lemma card_of_UNION_Sigma: -"|\i \ I. A i| \o |SIGMA i : I. A i|" -using Ex_inj_on_UNION_Sigma [of A I] card_of_ordLeq by blast + "|\i \ I. A i| \o |SIGMA i : I. A i|" + using Ex_inj_on_UNION_Sigma [of A I] card_of_ordLeq by blast 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" + assumes "a1 \ a2" + shows "|UNIV::bool set| =o |{a1,a2}|" +proof - + let ?f = "\ bl. if bl then a1 else a2" have "bij_betw ?f UNIV {a1,a2}" - proof- - {fix bl1 and bl2 assume "?f bl1 = ?f bl2" - hence "bl1 = bl2" using assms by (cases bl1, cases bl2) auto - } + proof - + have "\bl1 bl2. ?f bl1 = ?f bl2 \ bl1 = bl2" + using assms by (force split: if_split_asm) moreover - {fix bl have "?f bl \ {a1,a2}" by (cases 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 blast + have "\bl. ?f bl \ {a1,a2}" + using assms by (force split: if_split_asm) + ultimately show ?thesis unfolding bij_betw_def inj_on_def by force 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- + 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 blast - (* *) + using A2 card_of_mono1[of "{a1,a2}"] card_of_bool[of a1 a2] + by (blast intro: ordIso_ordLeq_trans) have "|A <+> B| \o |B <+> B|" - using LEQ card_of_Plus_mono1 by blast + 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 + 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) + using 1 by (simp add: card_of_Times_mono2) moreover have " |B \ A| =o |A \ B|" - using card_of_Times_commute by blast + 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 + by (blast intro: ordLeq_transitive dest: ordLeq_ordIso_trans) 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- + 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) + 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 blast + 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 blast } ultimately show ?thesis - using card_of_Well_order[of A] card_of_Well_order[of B] - ordLeq_total[of "|A|"] by blast + using card_of_Well_order[of A] card_of_Well_order[of B] + ordLeq_total[of "|A|"] by blast qed lemma card_of_Times_Plus_distrib: @@ -802,27 +729,27 @@ 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 + 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 "\ finite A" -shows "\ finite B" -using assms card_of_ordLeq_finite by auto + assumes "|A| \o |B|" and "\ finite A" + shows "\ finite 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) + 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 + 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\ @@ -835,18 +762,18 @@ at page 47 in @{cite "card-book"}. Then everything else follows fairly easily.\ lemma infinite_iff_card_of_nat: -"\ finite A \ ( |UNIV::nat set| \o |A| )" -unfolding infinite_iff_countable_subset card_of_ordLeq .. + "\ finite A \ ( |UNIV::nat set| \o |A| )" + unfolding 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: "\finite (Field r)" -shows "\ (\a. Field r = under r a)" + assumes CARD: "Card_order r" and INF: "\finite (Field r)" + shows "\ (\a. Field r = 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 + using CARD unfolding wo_rel_def card_order_on_def order_on_defs by auto fix a assume *: "Field r = under r a" show False proof(cases "a \ Field r") @@ -857,115 +784,113 @@ let ?r' = "Restr r (underS r a)" assume Case2: "a \ Field r" hence 1: "under r a = underS r a \ {a} \ a \ underS r a" - using 0 Refl_under_underS[of r a] underS_notIn[of a r] by blast + using 0 Refl_under_underS[of r a] underS_notIn[of a r] by blast have 2: "wo_rel.ofilter r (underS r a) \ underS r a < Field r" - using 0 wo_rel.underS_ofilter * 1 Case2 by fast + using 0 wo_rel.underS_ofilter * 1 Case2 by fast hence "?r' Well_order ?r'" - using 2 0 Field_Restr_ofilter[of r] Well_order_Restr[of r] by blast + using 2 0 Field_Restr_ofilter[of r] Well_order_Restr[of r] by blast ultimately have "|underS r a| f. bij_betw f (under r a) (underS r a)" - using infinite_imp_bij_betw[of "Field r" a] INF * 1 by auto - hence "|under r a| =o |underS r a|" using card_of_ordIso by blast + using infinite_imp_bij_betw[of "Field r" a] INF * 1 by auto + hence "|under r a| =o |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 "\finite (Field r)" -and a: "a \ Field r" -shows "\b \ Field r. a \ b \ (a,b) \ r" -proof- + assumes r: "Card_order r" and "\finite (Field r)" + and a: "a \ Field r" + shows "\b \ Field r. a \ b \ (a,b) \ r" +proof - have "Field r \ under r a" - using assms Card_order_infinite_not_under by blast + using assms Card_order_infinite_not_under by blast moreover have "under r a \ Field r" - using under_Field . - ultimately have "under r a < Field r" by blast - then obtain b where 1: "b \ Field r \ \ (b,a) \ r" - unfolding under_def by blast + using under_Field . + ultimately obtain b where b: "b \ Field r \ \ (b,a) \ r" + unfolding 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 + using b 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 + using a r unfolding card_order_on_def well_order_on_def linear_order_on_def + total_on_def by blast + thus ?thesis using b ba by auto qed theorem Card_order_Times_same_infinite: -assumes CO: "Card_order r" and INF: "\finite(Field r)" -shows "|Field r \ Field r| \o r" -proof- - obtain phi where phi_def: - "phi = (\r::'a rel. Card_order r \ \finite(Field r) \ - \ |Field r \ Field r| \o r )" by blast + assumes CO: "Card_order r" and INF: "\finite(Field r)" + shows "|Field r \ Field r| \o r" +proof - + define phi where + "phi \ \r::'a rel. Card_order r \ \finite(Field r) \ \ |Field r \ Field r| \o r" have temp1: "\r. phi r \ Well_order r" - unfolding phi_def card_order_on_def by auto + 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 + 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 + 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 + 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 - (* *) + 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 + 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 + 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) + 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 + 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 + 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 + 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 + using 4 wo_rel_def[of ?r'] wo_rel.ofilter_def[of ?r' ?B] by blast have "\ (\a. Field r = under r a)" - using 1 unfolding phi_def using Card_order_infinite_not_under[of r] by auto + 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 + 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 + hence "|A1| \o ?r1" using 3 Well_order_Restr card_of_least by blast } ultimately have 11: "|A1| finite (Field r)" using 1 unfolding phi_def by simp hence "\ finite ?B" using 8 3 card_of_ordIso_finite_Field[of r ?B] by blast hence "\ finite A1" using 9 finite_cartesian_product finite_subset by blast moreover have temp4: "Field |A1| = 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) + 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 + using temp4 11 3 using not_ordLeq_iff_ordLess by blast ultimately have "\ finite(Field |A1| ) \ Card_order |A1| \ \ r \o | A1 |" - by (simp add: card_of_card_order_on) + by (simp add: card_of_card_order_on) hence "|Field |A1| \ Field |A1| | \o |A1|" - using 2 unfolding phi_def by blast + 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 @@ -974,175 +899,175 @@ qed corollary card_of_Times_same_infinite: -assumes "\finite A" -shows "|A \ A| =o |A|" -proof- + assumes "\finite 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 + 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 + 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: "\finite A" and NE: "B \ {}" and LEQ: "|B| \o |A|" -shows "|A \ B| =o |A| \ |B \ A| =o |A|" -proof- + assumes INF: "\finite 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) + 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 + 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: "\finite(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- + assumes INF: "\finite(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) + 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 + using assms card_of_Field_ordIso by (blast intro: ordIso_transitive) qed lemma card_of_Sigma_ordLeq_infinite: -assumes INF: "\finite 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 \ {}" + assumes INF: "\finite 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 = {}") + case False have "|SIGMA i : I. A i| \o |I \ B|" - using card_of_Sigma_mono1[OF LEQ] by blast + using card_of_Sigma_mono1[OF LEQ] by blast moreover have "|I \ B| =o |B|" - using INF * LEQ_I by (auto simp add: card_of_Times_infinite) + using INF False LEQ_I by (auto simp add: card_of_Times_infinite) ultimately show ?thesis using ordLeq_ordIso_trans by blast -qed +qed (simp add: card_of_empty) lemma card_of_Sigma_ordLeq_infinite_Field: -assumes INF: "\finite (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- + assumes INF: "\finite (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 + 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+ + 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 + card_of_Sigma_ordLeq_infinite by blast thus ?thesis using 1 ordLeq_ordIso_trans by blast qed lemma card_of_Times_ordLeq_infinite_Field: -"\\finite (Field r); |A| \o r; |B| \o r; Card_order r\ - \ |A \ B| \o r" -by(simp add: card_of_Sigma_ordLeq_infinite_Field) + "\\finite (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: -"\\finite A; B \ {}; |B| \o |A|\ \ |A \ B| =o |A|" -"\\finite A; B \ {}; |B| \o |A|\ \ |A| =o |A \ B|" -"\\finite A; B \ {}; |B| \o |A|\ \ |B \ A| =o |A|" -"\\finite A; B \ {}; |B| \o |A|\ \ |A| =o |B \ A|" -by (auto simp add: card_of_Times_infinite ordIso_symmetric) + "\\finite A; B \ {}; |B| \o |A|\ \ |A \ B| =o |A|" + "\\finite A; B \ {}; |B| \o |A|\ \ |A| =o |A \ B|" + "\\finite A; B \ {}; |B| \o |A|\ \ |B \ A| =o |A|" + "\\finite 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: "\finite 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 \ {}" + assumes INF: "\finite 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 = {}") + case False have "|\i \ I. A i| \o |SIGMA i : I. A i|" - using card_of_UNION_Sigma by blast + 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 + using assms card_of_Sigma_ordLeq_infinite by blast ultimately show ?thesis using ordLeq_transitive by blast -qed +qed (simp add: card_of_empty) corollary card_of_UNION_ordLeq_infinite_Field: -assumes INF: "\finite (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- + assumes INF: "\finite (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 + 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+ + 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 + card_of_UNION_ordLeq_infinite by blast thus ?thesis using 1 ordLeq_ordIso_trans by blast qed lemma card_of_Plus_infinite1: -assumes INF: "\finite 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) + assumes INF: "\finite A" and LEQ: "|B| \o |A|" + shows "|A <+> B| =o |A|" +proof(cases "B = {}") + case True + then show ?thesis + by (simp add: card_of_Plus_empty1 card_of_Plus_empty2 ordIso_symmetric) +next + case False 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}" + case True have 2: "bij_betw ?Inl A ((?Inl ` A))" - unfolding bij_betw_def inj_on_def by auto + unfolding bij_betw_def inj_on_def by auto hence 3: "\finite (?Inl ` A)" - using INF bij_betw_finite[of ?Inl A] by blast + 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 + using 3 infinite_imp_bij_betw2[of "?Inl ` A"] by auto + moreover have "?A' = A <+> B" using True by blast ultimately have "bij_betw g (?Inl ` A) (A <+> B)" by simp hence "bij_betw (g \ ?Inl) A (A <+> B)" - using 2 by (auto simp add: bij_betw_trans) + using 2 by (auto simp add: bij_betw_trans) thus ?thesis using card_of_ordIso ordIso_symmetric by blast next - assume Case2: "B \ {b1}" + case False 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 + 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 + unfolding inj_on_def by auto with 3 have "|A <+> B| \o |A \ B|" - by (auto simp add: card_of_Plus_Times) + by (auto simp add: card_of_Plus_Times) moreover have "|A \ B| =o |A|" - using assms * by (simp add: card_of_Times_infinite_simps) + using assms * by (simp add: card_of_Times_infinite_simps) ultimately have "|A <+> B| \o |A|" using ordLeq_ordIso_trans by blast thus ?thesis using card_of_Plus1 ordIso_iff_ordLeq by blast qed qed lemma card_of_Plus_infinite2: -assumes INF: "\finite 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 + assumes INF: "\finite 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: "\finite 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) + assumes INF: "\finite 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: "\finite(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- + assumes INF: "\finite(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) + 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 + using assms card_of_Field_ordIso by (blast intro: ordIso_transitive) + qed @@ -1158,110 +1083,111 @@ definition "(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}" + where "natLeq_on n \ {(x,y). x < n \ y < n \ x \ y}" lemma infinite_cartesian_product: -assumes "\finite A" "\finite B" -shows "\finite (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 + assumes "\finite A" "\finite B" + shows "\finite (A \ B)" +using assms finite_cartesian_productD2 by auto subsubsection \First as well-orders\ lemma Field_natLeq: "Field natLeq = (UNIV::nat set)" -by(unfold Field_def natLeq_def, auto) + by(unfold Field_def natLeq_def, auto) lemma natLeq_Refl: "Refl natLeq" -unfolding refl_on_def Field_def natLeq_def by auto + unfolding refl_on_def Field_def natLeq_def by auto lemma natLeq_trans: "trans natLeq" -unfolding trans_def natLeq_def by auto + unfolding trans_def natLeq_def by auto lemma natLeq_Preorder: "Preorder natLeq" -unfolding preorder_on_def -by (auto simp add: natLeq_Refl natLeq_trans) + unfolding preorder_on_def + by (auto simp add: natLeq_Refl natLeq_trans) lemma natLeq_antisym: "antisym natLeq" -unfolding antisym_def natLeq_def by auto + unfolding antisym_def natLeq_def by auto lemma natLeq_Partial_order: "Partial_order natLeq" -unfolding partial_order_on_def -by (auto simp add: natLeq_Preorder natLeq_antisym) + unfolding partial_order_on_def + by (auto simp add: natLeq_Preorder natLeq_antisym) lemma natLeq_Total: "Total natLeq" -unfolding total_on_def natLeq_def by auto + unfolding total_on_def natLeq_def by auto lemma natLeq_Linear_order: "Linear_order natLeq" -unfolding linear_order_on_def -by (auto simp add: natLeq_Partial_order natLeq_Total) + unfolding linear_order_on_def + by (auto simp add: natLeq_Partial_order natLeq_Total) lemma natLeq_natLess_Id: "natLess = natLeq - Id" -unfolding natLeq_def natLess_def by auto + unfolding natLeq_def natLess_def by auto lemma natLeq_Well_order: "Well_order natLeq" -unfolding well_order_on_def -using natLeq_Linear_order wf_less natLeq_natLess_Id natLeq_def natLess_def by auto + unfolding well_order_on_def + using natLeq_Linear_order wf_less natLeq_natLess_Id natLeq_def natLess_def by auto lemma Field_natLeq_on: "Field (natLeq_on n) = {x. x < n}" -unfolding Field_def by auto + unfolding Field_def by auto lemma natLeq_underS_less: "underS natLeq n = {x. x < n}" -unfolding underS_def natLeq_def by auto + unfolding underS_def natLeq_def by auto lemma Restr_natLeq: "Restr natLeq {x. x < n} = natLeq_on n" -unfolding natLeq_def by force + unfolding natLeq_def by force lemma Restr_natLeq2: -"Restr natLeq (underS natLeq n) = natLeq_on n" -by (auto simp add: Restr_natLeq natLeq_underS_less) + "Restr natLeq (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 "{x. x < n}"] by auto + using Restr_natLeq[of n] natLeq_Well_order + Well_order_Restr[of natLeq "{x. x < n}"] by auto corollary natLeq_on_well_order_on: "well_order_on {x. x < n} (natLeq_on n)" -using natLeq_on_Well_order Field_natLeq_on by auto + using natLeq_on_Well_order Field_natLeq_on by auto lemma natLeq_on_wo_rel: "wo_rel(natLeq_on n)" -unfolding wo_rel_def using natLeq_on_Well_order . + unfolding wo_rel_def using natLeq_on_Well_order . 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))" by (auto simp: Field_def) - moreover have "\finite(UNIV::nat set)" by auto - ultimately show "natLeq_on n finite(UNIV::nat set)" by auto + ultimately show ?thesis + using finite_ordLess_infinite[of "natLeq_on n" "|UNIV::nat set|"] + card_of_Well_order[of "UNIV::nat set"] natLeq_on_Well_order + by (force simp add: Field_card_of) + qed + then show ?thesis + apply (simp add: natLeq_Well_order Card_order_iff_Restr_underS Restr_natLeq2) + apply (force simp add: Field_natLeq) + done qed corollary card_of_Field_natLeq: -"|Field natLeq| =o natLeq" -using Field_natLeq natLeq_Card_order Card_order_iff_ordIso_card_of[of natLeq] - ordIso_symmetric[of natLeq] by blast + "|Field natLeq| =o natLeq" + using Field_natLeq natLeq_Card_order Card_order_iff_ordIso_card_of[of natLeq] + ordIso_symmetric[of natLeq] by blast corollary card_of_nat: -"|UNIV::nat set| =o natLeq" -using Field_natLeq card_of_Field_natLeq by auto + "|UNIV::nat set| =o natLeq" + using Field_natLeq card_of_Field_natLeq by auto corollary infinite_iff_natLeq_ordLeq: -"\finite A = ( natLeq \o |A| )" -using infinite_iff_card_of_nat[of A] card_of_nat - ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast + "\finite A = ( natLeq \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| The successor of a cardinal\ @@ -1271,169 +1197,167 @@ not require \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'')" + 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 \cardSuc\, by picking {\em some} cardinal-order relation fulfilling \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'" + 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 + "\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- + 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) + 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 + 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) + 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 + "Card_order r \ Card_order(cardSuc r)" + using cardSuc_isCardSuc unfolding isCardSuc_def by blast lemma cardSuc_greater: -"Card_order r \ r r r \o cardSuc r" -using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast + "Card_order r \ r \o cardSuc r" + using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast text\The minimality property of \cardSuc\ originally present in its definition is local to the type \'a set rel\, i.e., that of \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 + "\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- + 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 + 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 + show "cardSuc r \o r' \ r o r)" -proof- + assumes CARD: "Card_order r" and CARD': "Card_order r'" + shows "(r' o r)" +proof - have "Well_order r \ Well_order r'" - using assms unfolding card_order_on_def by auto + 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 + 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 + using assms cardSuc_ordLess_ordLeq by (blast dest: not_ordLeq_iff_ordLess) 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 + 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- + 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) + 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 + 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 card_of_cardSuc_finite: -"finite(Field(cardSuc |A| )) = finite A" + "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 + 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 + 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 have "finite ( Field |Pow A| )" unfolding Field_card_of by simp - then show "finite (Field (cardSuc |A| ))" - proof (rule card_of_ordLeq_finite[OF card_of_mono2, rotated]) - show "cardSuc |A| \o |Pow A|" - by (rule iffD1[OF cardSuc_ordLess_ordLeq card_of_Pow]) (simp_all add: card_of_Card_order) - qed + moreover + have "cardSuc |A| \o |Pow A|" + by (rule iffD1[OF cardSuc_ordLess_ordLeq card_of_Pow]) (simp_all add: card_of_Card_order) + ultimately show "finite (Field (cardSuc |A| ))" + by (blast intro: card_of_ordLeq_finite card_of_mono2) qed lemma cardSuc_finite: -assumes "Card_order r" -shows "finite (Field (cardSuc r)) = finite (Field r)" -proof- + 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) + 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) + 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 + { 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 + using ordIso_transitive by blast hence "finite (Field (cardSuc |?A| )) = finite (Field (cardSuc r))" - using card_of_ordIso_finite by blast + using card_of_ordIso_finite by blast thus ?thesis by (simp only: card_of_cardSuc_finite) qed lemma Field_cardSuc_not_empty: -assumes "Card_order r" -shows "Field (cardSuc r) \ {}" + assumes "Card_order r" + shows "Field (cardSuc r) \ {}" proof assume "Field(cardSuc r) = {}" then have "|Field(cardSuc r)| \o r" using assms Card_order_empty[of r] by auto then have "cardSuc r \o r" using assms card_of_Field_ordIso - cardSuc_Card_order ordIso_symmetric ordIso_ordLeq_trans by blast + cardSuc_Card_order ordIso_symmetric ordIso_ordLeq_trans by blast then show False using cardSuc_greater not_ordLess_ordLeq assms by blast qed @@ -1487,265 +1411,263 @@ using ordLess_ordIso_trans[OF cardSuc_greater cardSuc_ordIso_card_suc] card_order_on_Card_order by blast lemma card_of_Plus_ordLess_infinite: -assumes INF: "\finite C" and - LESS1: "|A| B| finite C" and LESS1: "|A| B| B = {}") - assume Case1: "A = {} \ B = {}" + case True hence "|A| =o |A <+> B| \ |B| =o |A <+> B|" - using card_of_Plus_empty1 card_of_Plus_empty2 by blast + 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 + 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 + 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 "\finite (A <+> B)" using INF card_of_ordLeq_finite by blast - hence 1: "\finite A \ \finite B" using finite_Plus by blast - {assume Case21: "|A| \o |B|" - hence "\finite 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 "\finite 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 + case False + have False if "|C| \o |A <+> B|" + proof - + have \
: "\finite A \ \finite B" + using that INF card_of_ordLeq_finite finite_Plus by blast + consider "|A| \o |B|" | "|B| \o |A|" + using ordLeq_total [OF card_of_Well_order card_of_Well_order] by blast + then show False + proof cases + case 1 + hence "\finite B" using \
card_of_ordLeq_finite by blast + hence "|A <+> B| =o |B|" using False 1 + by (auto simp add: card_of_Plus_infinite) + thus False using LESS2 not_ordLess_ordLeq that ordLeq_ordIso_trans by blast + next + case 2 + hence "\finite A" using \
card_of_ordLeq_finite by blast + hence "|A <+> B| =o |A|" using False 2 + by (auto simp add: card_of_Plus_infinite) + thus False using LESS1 not_ordLess_ordLeq that ordLeq_ordIso_trans by blast + qed + qed + 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: "\finite (Field r)" and r: "Card_order r" and - LESS1: "|A| B| finite (Field r)" and r: "Card_order r" and + LESS1: "|A| B| |?C| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast + have 1: "r =o |?C| \ |?C| =o r" + using r card_of_Field_ordIso ordIso_symmetric by blast hence "|A| B| finite (Field r)" and A: "|A| \o r" and B: "|B| \o r" -and c: "Card_order r" -shows "|A <+> B| \o r" -proof- + assumes r: "\finite (Field r)" and A: "|A| \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' \ \finite (Field ?r')" using assms - by (simp add: cardSuc_Card_order cardSuc_finite) + by (simp add: cardSuc_Card_order cardSuc_finite) moreover have "|A| B| finite (Field r)" and A: "|A| \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 fast + assumes C: "\finite (Field r)" and A: "|A| \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 fast lemma card_of_Un_ordLess_infinite: -assumes INF: "\finite C" and - LESS1: "|A| B| finite C" and + LESS1: "|A| B| finite (Field r)" and r: "Card_order r" and - LESS1: "|A| finite (Field r)" and r: "Card_order r" and + LESS1: "|A| |?C| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast + ordIso_symmetric by blast hence "|A| Regular cardinals\ definition cofinal where -"cofinal A r \ - \a \ Field r. \b \ A. a \ b \ (a,b) \ r" + "cofinal A r \ \a \ Field r. \b \ A. a \ b \ (a,b) \ r" definition regularCard where -"regularCard r \ - \K. K \ Field r \ cofinal K r \ |K| =o r" + "regularCard r \ \K. K \ Field r \ cofinal K r \ |K| =o r" definition relChain where -"relChain r As \ - \i j. (i,j) \ r \ As i \ As j" + "relChain r As \ \i j. (i,j) \ r \ As i \ As j" lemma regularCard_UNION: -assumes r: "Card_order r" "regularCard r" -and As: "relChain r As" -and Bsub: "B \ (\i \ Field r. As i)" -and cardB: "|B| i \ Field r. B \ As i" -proof- + assumes r: "Card_order r" "regularCard r" + and As: "relChain r As" + and Bsub: "B \ (\i \ Field r. As i)" + and cardB: "|B| i \ Field r. B \ As i" +proof - let ?phi = "\b j. j \ Field r \ b \ As j" have "\b\B. \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 + 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 regularCard_def by blast - moreover - { - have "|?K| <=o |B|" using card_of_image . - hence "|?K| 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 \ f`B. i \ b \ (i,b) \ r" by blast + qed + moreover have "?K \ Field r" using f by blast + ultimately have "|?K| =o r" using 2 r unfolding regularCard_def by blast + moreover + have "|?K| finite (Field r)" and r_card: "Card_order r" -shows "regularCard (cardSuc r)" -proof- + assumes r_inf: "\finite (Field r)" and r_card: "Card_order r" + shows "regularCard (cardSuc r)" +proof - let ?r' = "cardSuc r" - have r': "Card_order ?r'" - "!! p. Card_order p \ (p \o r) = (p p. Card_order p \ (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']) + using r' by (simp add: card_of_Field_ordIso[of ?r']) finally have "|K| \o ?r'" . moreover - {let ?L = "\ j \ K. 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 "\j\K. |underS ?r' j| j\K. |underS ?r' j| \o r" - using r' card_of_Card_order by blast - hence "\j\K. |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 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| o r" using r' card_of_Card_order[of K] by blast + hence "|K| \o |?J|" using rJ ordLeq_ordIso_trans by blast + moreover + {have "\j\K. |underS ?r' j| j\K. |underS ?r' j| \o r" + using r' card_of_Card_order by blast + hence "\j\K. |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 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| finite (Field r)" -and As: "relChain (cardSuc r) As" -and Bsub: "B \ (\ i \ Field (cardSuc r). As i)" -and cardB: "|B| <=o r" -shows "\i \ Field (cardSuc r). B \ As i" -proof- + assumes r: "Card_order r" and "\finite (Field r)" + and As: "relChain (cardSuc r) As" + and Bsub: "B \ (\ i \ Field (cardSuc r). As i)" + and cardB: "|B| \o r" + shows "\i \ Field (cardSuc r). B \ As i" +proof - let ?r' = "cardSuc r" have "Card_order ?r' \ |B| Others\ 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 + "|Func (A \ B) C| =o |Func A (Func B C)|" + unfolding card_of_ordIso[symmetric] + using bij_betw_curr by blast lemma card_of_Pow_Func: -"|Pow A| =o |Func A (UNIV::bool set)|" -proof- - define F where [abs_def]: "F A' a = + "|Pow A| =o |Func A (UNIV::bool set)|" +proof - + define F where [abs_def]: "F A' a \ (if a \ A then (if a \ A' then True else False) else undefined)" for A' a have "bij_betw F (Pow A) (Func A (UNIV::bool set))" - unfolding bij_betw_def inj_on_def proof (intro ballI impI conjI) + 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: if_split_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 + show "f \ F ` Pow A" + unfolding image_iff + proof + show "f = F {a \ A. f a = True}" + using f unfolding Func_def F_def by force qed auto - qed(unfold Func_def mem_Collect_eq F_def, auto) + qed(unfold Func_def 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) + "|Func (UNIV::'a set) (B::'b set)| =o |{f::'a \ 'b. range f \ B}|" +proof - 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 + have "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 blast hence "range f \ B" using h unfolding Func_def by auto thus "h \ (\f a. f a) ` {f. range f \ B}" using f by auto qed(unfold Func_def fun_eq_iff, auto) + then show ?thesis + by (blast intro: ordIso_symmetric card_of_ordIsoI) qed lemma Func_Times_Range: @@ -1774,8 +1696,8 @@ subsection \Regular vs. stable cardinals\ definition stable :: "'a rel \ bool" -where -"stable r \ \(A::'a set) (F :: 'a \ 'a set). + where + "stable r \ \(A::'a set) (F :: 'a \ 'a set). |A| (\a \ A. |F a| |SIGMA a : A. F a| A" define L where "L = {(a,u) | u. u \ F a}" have fL: "f ` L \ Field r" using f a unfolding L_def by auto - have "|L| =o |F a|" unfolding L_def card_of_ordIso[symmetric] - apply(rule exI[of _ snd]) unfolding bij_betw_def inj_on_def by (auto simp: image_def) + have "bij_betw snd {(a, u) |u. u \ F a} (F a)" + unfolding bij_betw_def inj_on_def by (auto simp: image_def) + then have "|L| =o |F a|" unfolding L_def card_of_ordIso[symmetric] by blast hence "|L| cofinal (f ` L) r" using reg fL unfolding regularCard_def - apply simp - apply (drule not_ordLess_ordIso) - by auto + by (force simp add: dest: not_ordLess_ordIso) then obtain k where k: "k \ Field r" and "\ l \ L. \ (f l \ k \ (k, f l) \ r)" unfolding cofinal_def image_def by auto hence "\ k \ Field r. \ l \ L. (f l, k) \ r" @@ -1816,11 +1737,13 @@ hence 1: "Field r \ (\a \ A. under r (g a))" using f[symmetric] unfolding under_def image_def by auto have gA: "g ` A \ Field r" using gg j0 unfolding Field_def g_def by auto - moreover have "cofinal (g ` A) r" unfolding cofinal_def proof safe + moreover have "cofinal (g ` A) r" unfolding cofinal_def + proof safe fix i assume "i \ Field r" then obtain j where ij: "(i,j) \ r" "i \ j" using cr ir infinite_Card_order_limit by fast hence "j \ Field r" using card_order_on_def cr well_order_on_domain by fast - then obtain a where a: "a \ A" and j: "(j, g a) \ r" using 1 unfolding under_def by auto + then obtain a where a: "a \ A" and j: "(j, g a) \ r" + using 1 unfolding under_def by auto hence "(i, g a) \ r" using ij wo_rel.TRANS[OF r] unfolding trans_def by blast moreover have "i \ g a" using ij j r unfolding wo_rel_def unfolding well_order_on_def linear_order_on_def @@ -1836,40 +1759,40 @@ qed lemma stable_regularCard: -assumes cr: "Card_order r" and ir: "\finite (Field r)" and st: "stable r" -shows "regularCard r" -unfolding regularCard_def proof safe + assumes cr: "Card_order r" and ir: "\finite (Field r)" and st: "stable r" + shows "regularCard r" + unfolding regularCard_def proof safe fix K assume K: "K \ Field r" and cof: "cofinal K r" have "|K| \o r" using K card_of_Field_ordIso card_of_mono1 cr ordLeq_ordIso_trans by blast moreover {assume Kr: "|K| a. a \ Field r \ \b. b \ K \ a \ b \ (a, b) \ r" using cof unfolding cofinal_def by blast - then obtain f where "\a. a \ Field r \ f a = (SOME b. b \ K \ a \ b \ (a, b) \ r)" by simp - then have "\a\Field r. f a \ K \ a \ f a \ (a, f a) \ r" using someI_ex[OF x] by simp - hence "Field r \ (\a \ K. underS r a)" unfolding underS_def by auto - hence "r \o |\a \ K. underS r a|" - using cr Card_order_iff_ordLeq_card_of card_of_mono1 ordLeq_transitive by blast - also have "|\a \ K. underS r a| \o |SIGMA a: K. underS r a|" by (rule card_of_UNION_Sigma) - also - {have "\ a \ K. |underS r a| a. a \ Field r \ \b. b \ K \ a \ b \ (a, b) \ r" using cof unfolding cofinal_def by blast + then obtain f where "\a. a \ Field r \ f a = (SOME b. b \ K \ a \ b \ (a, b) \ r)" by simp + then have "\a\Field r. f a \ K \ a \ f a \ (a, f a) \ r" using someI_ex[OF x] by simp + hence "Field r \ (\a \ K. underS r a)" unfolding underS_def by auto + hence "r \o |\a \ K. underS r a|" + using cr Card_order_iff_ordLeq_card_of card_of_mono1 ordLeq_transitive by blast + also have "|\a \ K. underS r a| \o |SIGMA a: K. underS r a|" by (rule card_of_UNION_Sigma) + also + {have "\ a \ K. |underS r a| B < Field r. |A| =o |B| \ |B| B < Field r. |A| =o |B| \ |B| |A| =o p \ p |Field p| B < Field r. |A| =o |B| \ |B| B < Field r. |A| =o |B| \ |B| i \ I. |A i| =o |B i|" -shows "|SIGMA i : I. A i| =o |SIGMA i : I. B i|" -using assms by (auto simp add: card_of_Sigma_mono1 ordIso_iff_ordLeq) + assumes "\i \ I. |A i| =o |B i|" + shows "|SIGMA i : I. A i| =o |SIGMA i : I. B i|" + using assms by (auto simp add: card_of_Sigma_mono1 ordIso_iff_ordLeq) lemma card_of_Sigma_cong2: -assumes "bij_betw f (I::'i set) (J::'j set)" -shows "|SIGMA i : I. (A::'j \ 'a set) (f i)| =o |SIGMA j : J. A j|" -proof- + assumes "bij_betw f (I::'i set) (J::'j set)" + shows "|SIGMA i : I. (A::'j \ 'a set) (f i)| =o |SIGMA j : J. A j|" +proof - let ?LEFT = "SIGMA i : I. A (f i)" let ?RIGHT = "SIGMA j : J. A j" - obtain u where u_def: "u = (\(i::'i,a::'a). (f i,a))" by blast + define u where "u \ \(i::'i,a::'a). (f i,a)" have "bij_betw u ?LEFT ?RIGHT" - using assms unfolding u_def bij_betw_def inj_on_def by auto + using assms unfolding u_def bij_betw_def inj_on_def by auto thus ?thesis using card_of_ordIso by blast qed lemma card_of_Sigma_cong: -assumes BIJ: "bij_betw f I J" and - ISO: "\j \ J. |A j| =o |B j|" -shows "|SIGMA i : I. A (f i)| =o |SIGMA j : J. B j|" -proof- + assumes BIJ: "bij_betw f I J" and ISO: "\j \ J. |A j| =o |B j|" + shows "|SIGMA i : I. A (f i)| =o |SIGMA j : J. B j|" +proof - have "\i \ I. |A(f i)| =o |B(f i)|" - using ISO BIJ unfolding bij_betw_def by blast + using ISO BIJ unfolding bij_betw_def by blast hence "|SIGMA i : I. A (f i)| =o |SIGMA i : I. B (f i)|" by (rule card_of_Sigma_cong1) moreover have "|SIGMA i : I. B (f i)| =o |SIGMA j : J. B j|" - using BIJ card_of_Sigma_cong2 by blast + using BIJ card_of_Sigma_cong2 by blast ultimately show ?thesis using ordIso_transitive by blast qed (* Note that below the types of A and F are now unconstrained: *) lemma stable_elim: -assumes ST: "stable r" and A_LESS: "|A| a. a \ A \ |F a| a. a \ A \ |F a| |A'| A" - hence "\B'. B' \ Field r \ |F a| =o |B'| \ |B'| B'. B' \ Field r \ |F a| =o |B'| \ |B'| a \ A. F' a \ Field r \ |F a| =o |F' a| \ |F' a| a B'. B' \ Field r \ |F a| =o |B'| \ |B'| a \ A. F' a \ Field r \ |F a| =o |F' a| \ |F' a| a B'. B' \ Field r \ |F a| =o |B'| \ |B'| a \ A. F' a \ Field r \ |F' a| a \ A. |F' a| =o |F a|" using temp ordIso_symmetric by auto - (* *) + (* *) have "\a' \ A'. F'(G a') \ Field r \ |F'(G a')| 'a set" assume "|A| a\A. |F a| (\a \ A. finite(F a))" - by (auto simp add: finite_iff_ordLess_natLeq) + by (auto simp add: finite_iff_ordLess_natLeq) hence "finite(Sigma A F)" by (simp only: finite_SigmaI[of A F]) thus "|Sigma A F | 'b set" assume "|A| a \ A. |F a| (\a \ A. |F a| a. a \ A \ |F a| a \ A. F a| a \ A. F a| \o |SIGMA a : A. F a|" - using card_of_UNION_Sigma by blast - thus ?thesis using assms stable_elim ordLeq_ordLess_trans by blast -qed + assumes "stable r" and "|A| a. a \ A \ |F a| a \ A. F a| i \ I. |A i| i \ I. A i| i \ I. |A i| i \ I. A i| i \ I. |A i| i \ I. A i| i \ I. |A i| i \ I. A i| |?B| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast + ordIso_symmetric by blast hence "|I| i \ I. |A i| i \ I. A i| Constructions on Wellorders as Needed by Bounded Natural Functors\ theory BNF_Wellorder_Constructions -imports BNF_Wellorder_Embedding + imports BNF_Wellorder_Embedding begin text \In this section, we study basic constructions on well-orders, such as restriction to @@ -70,13 +70,8 @@ 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 + using assms + by (auto simp: well_order_on_def Linear_order_Restr elim: wf_subset) lemma Field_Restr_subset: "Field(Restr r A) \ A" by (auto simp add: Field_def) @@ -92,7 +87,7 @@ 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 assms using Well_order_Restr[of r A] Refl_Field_Restr2[of r A] order_on_defs[of "Field r" r] by auto @@ -100,12 +95,12 @@ 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) + "\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 "under (Restr r A) a = under r a" + assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and IN: "a \ A" + shows "under (Restr r A) a = under r a" unfolding wo_rel.ofilter_def under_def proof show "{b. (b, a) \ Restr r A} \ {b. (b, a) \ r}" @@ -125,51 +120,51 @@ qed lemma ofilter_embed: -assumes "Well_order r" -shows "wo_rel.ofilter r A = (A \ Field r \ embed (Restr r A) r id)" + 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" - unfolding embed_def + unfolding embed_def proof safe fix a assume "a \ A" thus "a \ Field r" using assms * - by (auto simp add: wo_rel_def wo_rel.ofilter_def) + by (auto simp add: wo_rel_def wo_rel.ofilter_def) next fix a assume "a \ Field (Restr r A)" thus "bij_betw id (under (Restr r A) a) (under r (id a))" using assms * - by (simp add: ofilter_Restr_under Field_Restr_ofilter) + 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 + 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 (under (Restr r A) a) (under r a)" - using * unfolding embed_def by auto - hence "under r a \ under (Restr r A) a" - unfolding bij_betw_def by auto - also have "\ \ Field(Restr r A)" by (simp add: under_Field) - also have "\ \ A" by (simp add: Field_Restr_subset) - finally have "under r a \ A" . + hence "a \ Field(Restr r A)" using * assms + by (simp add: order_on_defs Refl_Field_Restr2) + hence "bij_betw id (under (Restr r A) a) (under r a)" + using * unfolding embed_def by auto + hence "under r a \ under (Restr r A) a" + unfolding bij_betw_def by auto + also have "\ \ Field(Restr r A)" by (simp add: under_Field) + also have "\ \ A" by (simp add: Field_Restr_subset) + finally have "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)" + 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 + 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 *) + by (simp add: Well_order_Restr wo_rel_def) + (* Main proof *) show ?thesis using WellB assms unfolding wo_rel.ofilter_def under_def ofilter_def proof safe @@ -183,84 +178,84 @@ 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" + 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)" + 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 + using Refl_Field_Restr by blast have FieldB: "Field ?rB = Field r Int B" - using Refl Refl_Field_Restr by blast + 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) + 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 *) + 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) + by (simp add: ofilter_Restr_subset) hence "embed (Restr ?rB A) (Restr r B) id" - using WellB ofilter_embed[of "?rB" A] by auto + using WellB ofilter_embed[of "?rB" A] by auto thus "embed (Restr r A) (Restr r B) id" - using * by (simp add: Restr_subset) + 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 + 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)) \ + 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 + using Refl_Field_Restr by blast hence FieldA: "Field ?rA = A" using OFA Well - by (auto simp add: wo_rel.ofilter_def) + by (auto simp add: wo_rel.ofilter_def) have "Field ?rB = Field r Int B" - using Refl Refl_Field_Restr by blast + 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 *) + 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 + 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) + 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" + 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')" @@ -268,13 +263,13 @@ 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 + 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 + 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 @@ -283,13 +278,13 @@ 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 \ + 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)" + 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) @@ -299,23 +294,23 @@ moreover have "compat (ofilterIncl r) ?rS ?h" proof(unfold compat_def ofilterIncl_def, - intro allI impI, simp, elim conjE) + 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" + **: "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 = underS r a \ B = underS r b" - using Well by (auto simp add: wo_rel.ofilter_underS_Field) + 1: "A = underS r a \ B = 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: underS_incl_iff) + by (auto simp add: 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) + 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 + by simp qed ultimately show "wf (ofilterIncl r)" by (simp add: compat_wf) qed @@ -336,35 +331,35 @@ \ definition ordLeq :: "('a rel * 'a' rel) set" -where -"ordLeq = {(r,r'). Well_order r \ Well_order r' \ (\f. embed r r' f)}" + 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" + where "r <=o r' \ (r,r') \ ordLeq" abbreviation ordLeq3 :: "'a rel \ 'a' rel \ bool" (infix "\o" 50) -where "r \o r' \ r <=o r'" + 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)}" + 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" + where "r (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)}" + 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" + 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 + assumes "r \o r'" + shows "Well_order r \ Well_order r'" + using assms unfolding ordLeq_def by simp text\Notice that the relations \\o\, \, \=o\ connect well-orders on potentially {\em distinct} types. However, some of the lemmas below, including the next one, @@ -372,360 +367,313 @@ to \'a rel rel\.\ lemma ordLeq_reflexive: -"Well_order r \ r \o r" -unfolding ordLeq_def using id_embed[of r] by blast + "Well_order r \ 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' \ f)" - using comp_embed[of r r' f r'' f'] by auto - thus "r \o r''" unfolding ordLeq_def using 1 by auto -qed + assumes "r \o r'" and "r' \o r''" + shows "r \o r''" + using assms by (auto simp: ordLeq_def intro: comp_embed) lemma ordLeq_total: -"\Well_order r; Well_order r'\ \ r \o r' \ r' \o r" -unfolding ordLeq_def using wellorders_totally_ordered by blast + "\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 + "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' \ f)" - using comp_iso[of r r' f r'' f'] by auto - thus "r =o r''" unfolding ordIso_def using 1 by auto -qed + assumes *: "r =o r'" and **: "r' =o r''" + shows "r =o r''" + using assms by (auto simp: ordIso_def intro: comp_iso) lemma ordIso_symmetric: -assumes *: "r =o r'" -shows "r' =o r" + 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) + 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) + 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' o r'" and " r' Well_order r''" - using assms unfolding ordLeq_def ordLess_def by auto + using assms unfolding ordLeq_def ordLess_def by auto thus ?thesis using assms unfolding ordLeq_def ordLess_def - using embed_comp_embedS by blast + 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 + assumes "r o r''" + shows "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 + assumes "r \o r'" and " r' =o r''" + shows "r \o r''" + using embed_comp_iso assms by (force simp: ordLeq_def ordIso_def) 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 + assumes "r =o r'" and " r' \o r''" + shows "r \o r''" + using iso_comp_embed assms by (force simp: ordLeq_def ordIso_def) 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 + assumes "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 + assumes "r =o r'" and " r' (\f'. embed r' r f')" + 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) + 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 + 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)" + 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) + 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 + using 0 EMB embed_unique[of r1] by auto hence "\(bij_betw f ?A1 ?A2)" - using 1 bij_betw_cong[of ?A1] by blast + 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'))" + "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 + 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 + 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 + 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) + 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 + using id_embed[of r] by (auto simp: ordLess_iff) lemma ordLeq_iff_ordLess_or_ordIso: -"r \o r' = (r r =o r')" -unfolding ordRels_def embedS_defs iso_defs by blast + "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)" + "(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 + 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) + by (simp add: inv_into_Field_embed_bij_betw) thus "r \o r' \ r' \o r" - unfolding ordLeq_def using 1 by auto + 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 + 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 + "r \ r' \o r" + using ordLess_ordLeq_trans ordLess_irreflexive by blast lemma not_ordLeq_ordLess: -"r \o r' \ \ r' o r' \ \ r' r' \o r" + 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) + 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 + 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 ordLess_ordIso_trans ordIso_symmetric ordLess_irreflexive by blast + "r \ r =o r'" + using 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) = (r r' o r')" -using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "(\ r' o r')" + using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast lemma ordLess_transitive[trans]: -"\r \ r r \ r r \o r'" -using ordIso_iff_ordLeq by blast + "r =o 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 + "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)" + 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 + 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 + 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 + 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 + 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 + 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 + 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 Well_order r; wo_rel.ofilter r A\ \ (A < Field r) = (Restr r A {}" -shows "Restr r (underS r a) {}" + shows "Restr r (underS r a) (ofilterIncl r3)" + assumes + OL12: "r1 (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) + 0: "Well_order 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 + 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 - (* *) + 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) + 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) + 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) - (* *) + 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) + 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) + using EMB23 0 by (simp add: wo_rel_def embed_inj_on) ultimately have "f23 ` (f12 ` ?A1) < f23 ` ?A2" by (simp add: image_strict_mono) moreover {have "embed r1 r3 (f23 \ 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 \ f12" f13] by auto - hence "f23 ` (f12 ` ?A1) = f13 ` ?A1" by force + 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 \ 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 + 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 (underS r a))" + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "(r' a \ Field r. r' =o Restr r (underS r a))" proof safe fix a assume *: "a \ Field r" and **: "r' =o Restr r (underS r a)" hence "Restr r (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 + 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: "underS r a = f ` (Field r')" - using 1 2 by (auto simp add: wo_rel.ofilter_underS_Field wo_rel_def) + 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 + 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 + using WELL Well_order_Restr by blast ultimately have "r' =o Restr r (f ` (Field r'))" - using WELL' unfolding ordIso_def by auto + using WELL' unfolding ordIso_def by auto hence "r' =o Restr r (underS r a)" using 4 by auto thus "\a \ Field r. r' =o Restr r (underS r a)" using 3 by auto qed lemma internalize_ordLess: -"(r' p. Field p < Field r \ r' =o p \ p 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 (underS r a)" - using ordLess_iff_ordIso_Restr by blast + using ordLess_iff_ordIso_Restr by blast let ?p = "Restr r (underS r a)" have "wo_rel.ofilter r (underS r a)" using 0 - by (simp add: wo_rel_def wo_rel.underS_ofilter) + by (simp add: wo_rel_def wo_rel.underS_ofilter) hence "Field ?p = underS r a" using 0 Field_Restr_ofilter by blast hence "Field ?p < Field r" using underS_Field2 1 by fast moreover have "?p p. Field p < Field r \ r' =o p \ 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)" + "(r' \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 - } + have "r' \p. Field p \ Field r \ r' =o p \ p \o r" + using ordLeq_iff_ordLess_or_ordIso internalize_ordLess[of r' r] 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 + 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 (underS r a) o r') = (\a \ Field r. Restr r (underS r a) o r'" fix a assume "a \ Field r" hence "Restr r (underS r a) a \ Field r. Restr r (underS r a) Field r \ r' =o Restr r (underS r a)" - using assms ordLess_iff_ordIso_Restr by blast - hence False using * not_ordLess_ordIso ordIso_symmetric by blast + then obtain a where "a \ Field r \ r' =o Restr r (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: "\finite(Field r')" -shows "r finite(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 blast + 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 blast } 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'" + 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 well_order_on_Well_order by blast + using assms 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 * ** well_order_on_Well_order by blast + using * ** 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 + "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 @@ -854,10 +797,10 @@ {\em strict inclusion}; and we already know that strict inclusion of order filters is well-founded.\ definition ord_to_filter :: "'a rel \ 'a rel \ 'a set" -where "ord_to_filter r0 r \ (SOME f. embed r r0 f) ` (Field r)" + where "ord_to_filter r0 r \ (SOME f. embed r r0 f) ` (Field r)" lemma ord_to_filter_compat: -"compat (ordLess Int (ordLess\``{r0} \ ordLess\``{r0})) + "compat (ordLess Int (ordLess\``{r0} \ ordLess\``{r0})) (ofilterIncl r0) (ord_to_filter r0)" proof(unfold compat_def ord_to_filter_def, clarify) @@ -867,41 +810,41 @@ 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) + 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) + 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\``{r0} \ ?ordLess\``{r0})" - {assume Case1: "Well_order r0" - hence "wf ?R" - using wf_ofilterIncl[of r0] + (* need to annotate here!*) + let ?ordLess = "ordLess::('d rel * 'd rel) set" + let ?R = "?ordLess Int (?ordLess\``{r0} \ ?ordLess\``{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 + } + 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'" + 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' definition dir_image :: "'a rel \ ('a \ 'a') \ 'a' rel" -where -"dir_image r f = {(f a, f b)| a b. (a,b) \ r}" + 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 Range_def Domain_def by fast + "Field(dir_image r f) = f ` (Field r)" + unfolding dir_image_def Field_def Range_def Domain_def by fast 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 + "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)" + 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 + 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) + 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)" -unfolding trans_def + assumes TRANS: "trans r" and INJ: "inj_on f (Field r)" + shows "trans(dir_image r f)" + unfolding trans_def proof safe 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 + 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 + 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 + 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) + "\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)" -unfolding antisym_def + assumes AN: "antisym r" and INJ: "inj_on f (Field r)" + shows "antisym(dir_image r f)" + unfolding antisym_def proof safe 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 + 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) + "\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)" + 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)" @@ -995,76 +938,76 @@ 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 + 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) + "\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)" + 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 SUB NE by (auto simp: dir_image_Field) then obtain a where 1: "a \ A \ (\b \ A. (b,a) \ r)" - using spec[OF WF[unfolded wf_eq_minimal2], of A] by blast + using spec[OF WF[unfolded wf_eq_minimal2], of A] by blast 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 + 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 + 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)" -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 + "\Well_order r; inj_on f (Field r)\ \ Well_order (dir_image r f)" + 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_bij_betw: -"\inj_on f (Field r)\ \ bij_betw f (Field r) (Field (dir_image r f))" -unfolding bij_betw_def by (simp add: dir_image_Field order_on_defs) + "\inj_on f (Field r)\ \ bij_betw f (Field r) (Field (dir_image r f))" + unfolding bij_betw_def by (simp add: dir_image_Field order_on_defs) lemma dir_image_compat: -"compat r (dir_image r f) f" -unfolding compat_def dir_image_def by auto + "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 + "\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 + "\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'" + 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 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 + let ?r' = "dir_image r f" + have 1: "A = Field r \ Well_order r" + using WELL 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 @@ -1086,8 +1029,8 @@ 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)). + 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 \ @@ -1096,15 +1039,15 @@ )}" lemma Field_bsqr: -"Field (bsqr r) = Field r \ Field r" + "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 \ + 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 + ultimately have "a1 \ Field r \ a2 \ Field r" unfolding Field_def by auto } thus ?thesis unfolding Field_def by force qed @@ -1118,101 +1061,101 @@ qed lemma bsqr_Refl: "Refl(bsqr r)" -by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def) + by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def) lemma bsqr_Trans: -assumes "Well_order r" -shows "trans (bsqr r)" -unfolding trans_def + assumes "Well_order r" + shows "trans (bsqr r)" + unfolding trans_def proof safe (* 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 *) + (* 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 + 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 + 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 + 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 + {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 + {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 + {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)" + 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 @@ -1220,334 +1163,293 @@ 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 *) + 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 \ + hence "{a1,a2,b1,b2} \ Field r" unfolding bsqr_def by auto + moreover + have "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 \ + using * unfolding bsqr_def by auto + moreover + have "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 + using ** unfolding bsqr_def by auto + ultimately show "a1 = b1 \ a2 = b2" + using IrrS by auto qed lemma bsqr_Total: -assumes "Well_order r" -shows "Total(bsqr r)" + 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 *) + 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 + 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 + 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 + 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 + 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 + 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 - (* *) + 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 - (* *) + 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 + 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 - (* *) + 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 + 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 + 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" + 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 + 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 + 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 + 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 = under r a)" -shows "\A. wo_rel.ofilter r A \ A < Field r \ D \ A \ A" + 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 = 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 - (* *) + 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 = underS ?r' (a1,a2)" - using Well' wo_rel.ofilter_underS_Field[of ?r' D] by auto + "(a1,a2) \ Field ?r'" and 1: "D = 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 \ (under r ?m) \ (under r ?m)" proof(unfold 1) {fix b1 b2 - let ?n = "wo_rel.max2 r b1 b2" - assume "(b1,b2) \ underS ?r' (a1,a2)" - hence 3: "((b1,b2),(a1,a2)) \ ?r'" - unfolding 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) \ (under r ?m) \ (under r ?m)" unfolding under_def by simp} - thus "underS ?r' (a1,a2) \ (under r ?m) \ (under r ?m)" by auto + let ?n = "wo_rel.max2 r b1 b2" + assume "(b1,b2) \ underS ?r' (a1,a2)" + hence 3: "((b1,b2),(a1,a2)) \ ?r'" + unfolding 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) \ (under r ?m) \ (under r ?m)" unfolding under_def by simp} + thus "underS ?r' (a1,a2) \ (under r ?m) \ (under r ?m)" by auto qed moreover have "wo_rel.ofilter r (under r ?m)" - using Well by (simp add: wo_rel.under_ofilter) + using Well by (simp add: wo_rel.under_ofilter) moreover have "under r ?m < Field r" - using NE under_Field[of r ?m] by blast + using NE under_Field[of r ?m] by blast ultimately show ?thesis by blast qed definition Func where -"Func A B = {f . (\ a \ A. f a \ B) \ (\ a. a \ A \ f a = undefined)}" + "Func A B = {f . (\ a \ A. f a \ B) \ (\ a. a \ A \ f a = undefined)}" lemma Func_empty: -"Func {} B = {\x. undefined}" -unfolding Func_def by auto + "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 + 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" + "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 + 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" + 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" @@ -1559,14 +1461,14 @@ 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 + 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" + 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" @@ -1579,7 +1481,7 @@ next case True obtain g1 where "g1 \ Func B C" and "g a = g1" - using assms using Func_elim[OF assms True] by blast + using assms using Func_elim[OF assms True] by blast thus ?thesis using True unfolding Func_def curr_def by auto qed qed @@ -1587,24 +1489,21 @@ 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 - apply (intro impI conjI ballI) - apply (erule curr_inj[THEN iffD1], assumption+, safe) - using curr_surj curr_in apply blast+ - done + "bij_betw (curr A) (Func (A \ B) C) (Func A (Func B C))" + unfolding bij_betw_def inj_on_def + using curr_surj curr_in curr_inj by blast definition Func_map where -"Func_map B2 f1 f2 g b2 \ if b2 \ B2 then f1 (g (f2 b2)) else undefined" + "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 + 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 \ {}" + 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 @@ -1612,27 +1511,21 @@ qed lemma Func_is_emp: -"Func A B = {} \ A \ {} \ B = {}" (is "?L \ ?R") + "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[of B A] by simp } - ultimately show ?R by blast + assume ?L + then show ?R + using Func_empty Func_non_emp[of B A] 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 blast - with R have False by blast - } - thus ?L by blast + assume ?R + then show ?L + using Func_empty Func_non_emp[of B A] by (auto simp: Func_def) 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" + 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) @@ -1646,15 +1539,15 @@ then obtain k where k: "\ a2 \ f2 ` B2. k a2 \ B2 \ f2 (k a2) = a2" by atomize_elim (rule bchoice) {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 + 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 define j2 where [abs_def]: "j2 a2 = (if a2 \ f2 ` B2 then k a2 else b22)" for a2 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 + using kk unfolding j2_def by auto define g where "g = Func_map A2 j1 j2 h" have "Func_map B2 f1 f2 g = h" proof (rule ext) @@ -1668,13 +1561,13 @@ 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) + 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 j2A2 B1 A2 unfolding j1_def by (fast intro: inv_into_into)+ + using j2A2 B1 A2 unfolding j1_def by (fast intro: inv_into_into)+ ultimately show "h \ Func_map B2 f1 f2 ` Func A2 A1" - unfolding Func_map_def[abs_def] by auto + unfolding Func_map_def[abs_def] by auto qed(use B1 Func_map[OF _ _ A2(2)] in auto) qed diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/BNF_Wellorder_Embedding.thy --- a/src/HOL/BNF_Wellorder_Embedding.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/BNF_Wellorder_Embedding.thy Sun Jan 15 20:00:44 2023 +0100 @@ -8,7 +8,7 @@ section \Well-Order Embeddings as Needed by Bounded Natural Functors\ theory BNF_Wellorder_Embedding -imports Hilbert_Choice BNF_Wellorder_Relation + imports Hilbert_Choice BNF_Wellorder_Relation begin text\In this section, we introduce well-order {\em embeddings} and {\em isomorphisms} and @@ -23,35 +23,35 @@ 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)" + 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 + using wo_rel.ofilter_linord[of r] OF by blast with WELL INJ show ?thesis - by (auto simp add: inj_on_UNION_chain) + 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 (underS r a) (underS r' (f a))" -shows "bij_betw f (under r a) (under r' (f a))" + 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 (underS r a) (underS r' (f a))" + shows "bij_betw f (under r a) (under r' (f a))" proof- have "a \ underS r a \ f a \ underS r' (f a)" - unfolding underS_def by auto + unfolding underS_def by auto moreover {have "Refl r \ Refl r'" using WELL WELL' - by (auto simp add: order_on_defs) - hence "under r a = underS r a \ {a} \ + by (auto simp add: order_on_defs) + hence "under r a = underS r a \ {a} \ under r' (f a) = underS r' (f a) \ {f a}" - using IN IN' by(auto simp add: Refl_under_underS) + using IN IN' by(auto simp add: Refl_under_underS) } ultimately show ?thesis - using BIJ notIn_Un_bij_betw[of a "underS r a" f "underS r' (f a)"] by auto + using BIJ notIn_Un_bij_betw[of a "underS r a" f "underS r' (f a)"] by auto qed @@ -69,28 +69,28 @@ and an isomorphism (operator \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 (under r a) (under r' (f a))" + where + "embed r r' f \ \a \ Field r. bij_betw f (under r a) (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')" + 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')" + 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'" + 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'" @@ -168,46 +168,46 @@ 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)" + 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 *) + (* 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' \ under r' (f a)" hence "a \ Field r" using 0 by auto hence "bij_betw f (under r a) (under r' (f a))" - using * EMB by (auto simp add: embed_def) + using * EMB by (auto simp add: embed_def) hence "f`(under r a) = under r' (f a)" - by (simp add: bij_betw_def) + by (simp add: bij_betw_def) with ** image_def[of f "under r a"] obtain b where - 1: "b \ under r a \ b' = f b" by blast + 1: "b \ under r a \ b' = f b" by blast hence "b \ A" using Well * OF - by (auto simp add: wo_rel.ofilter_def) + 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))" + 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) + 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_inj_on: -assumes WELL: "Well_order r" and EMB: "embed r r' f" -shows "inj_on f (Field r)" + 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 . @@ -215,30 +215,30 @@ have Total: "Total r" by simp from Well wo_rel.REFL[of r] have Refl: "Refl r" by simp - (* Main proof *) + (* Main proof *) fix a b assume *: "a \ Field r" and **: "b \ Field r" and - ***: "f a = f b" + ***: "f a = f b" hence 1: "a \ Field r \ b \ Field r" - unfolding Field_def by auto + unfolding Field_def by auto {assume "(a,b) \ r" - hence "a \ under r b \ b \ under r b" - using Refl by(auto simp add: under_def refl_on_def) - hence "a = b" - using EMB 1 *** - by (auto simp add: embed_def bij_betw_def inj_on_def) + hence "a \ under r b \ b \ under r b" + using Refl by(auto simp add: 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 \ under r a \ b \ under r a" - using Refl by(auto simp add: under_def refl_on_def) - hence "a = b" - using EMB 1 *** - by (auto simp add: embed_def bij_betw_def inj_on_def) + hence "a \ under r a \ b \ under r a" + using Refl by(auto simp add: 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) + by (auto simp add: total_on_def) qed lemma embed_underS: @@ -265,173 +265,173 @@ 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 + 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 *) + 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 *) + **: "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 + using WELL wo_rel_def[of r] by simp hence Refl: "Refl r" - using wo_rel.REFL[of r] by simp + using wo_rel.REFL[of r] by simp have Total: "Total r" - using Well wo_rel.TOTAL[of r] by simp + using Well wo_rel.TOTAL[of r] by simp have Well': "wo_rel r'" - using WELL' wo_rel_def[of r'] by simp + using WELL' wo_rel_def[of r'] by simp hence Antisym': "antisym r'" - using wo_rel.ANTISYM[of r'] by simp + 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 + 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) + using ** by(auto simp add: compat_def) hence 0: "f a \ Field r'" - unfolding Field_def by auto + unfolding Field_def by auto have "f a \ f`(Field r)" - using **** by auto + using **** by auto hence 2: "under r' (f a) \ f`(Field r)" - using Well' *** wo_rel.ofilter_def[of r' "f`(Field r)"] by fastforce - (* Main proof *) + using Well' *** wo_rel.ofilter_def[of r' "f`(Field r)"] by fastforce + (* Main proof *) show "bij_betw f (under r a) (under r' (f a))" proof(unfold bij_betw_def, auto) show "inj_on f (under r a)" by (rule subset_inj_on[OF * under_Field]) next fix b assume "b \ under r a" thus "f b \ under r' (f a)" - unfolding under_def using ** - by (auto simp add: compat_def) + unfolding under_def using ** + by (auto simp add: compat_def) next fix b' assume *****: "b' \ under r' (f a)" hence "b' \ f`(Field r)" - using 2 by auto + using 2 by auto with Field_def[of r] obtain b where - 3: "b \ Field r" and 4: "b' = f b" by auto + 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: under_def antisym_def) - with 3 **** 4 * have "a = b" - by(auto simp add: inj_on_def) + with ** 4 have "(f a, b') \ r'" + by (auto simp add: compat_def) + with ***** Antisym' have "f a = b'" + by(auto simp add: 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) + 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`(under r a)" - unfolding under_def by auto + unfolding 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 (under r b) (under r' (f b))" and - IMAGE: "f ` A = Field r'" -shows "embed r' r (inv_into A f)" + assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and + BIJ: "\b \ A. bij_betw f (under r b) (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 + using WELL wo_rel_def[of r] by simp have Refl: "Refl r" - using Well wo_rel.REFL[of r] by simp + using Well wo_rel.REFL[of r] by simp have Total: "Total r" - using Well wo_rel.TOTAL[of r] by simp - (* Main proof *) + 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" + ***: "f b1 = f b2" have 11: "b1 \ Field r \ b2 \ Field r" - using * ** Well OF by (auto simp add: wo_rel.ofilter_def) + using * ** Well OF by (auto simp add: wo_rel.ofilter_def) moreover {assume "(b1,b2) \ r" - hence "b1 \ under r b2 \ b2 \ under r b2" - unfolding under_def using 11 Refl - by (auto simp add: refl_on_def) - hence "b1 = b2" using BIJ * ** *** - by (simp add: bij_betw_def inj_on_def) + hence "b1 \ under r b2 \ b2 \ under r b2" + unfolding under_def using 11 Refl + by (auto simp add: refl_on_def) + hence "b1 = b2" using BIJ * ** *** + by (simp add: bij_betw_def inj_on_def) } moreover - {assume "(b2,b1) \ r" - hence "b1 \ under r b1 \ b2 \ under r b1" - unfolding under_def using 11 Refl - by (auto simp add: refl_on_def) - hence "b1 = b2" using BIJ * ** *** - by (simp add: bij_betw_def inj_on_def) + {assume "(b2,b1) \ r" + hence "b1 \ under r b1 \ b2 \ under r b1" + unfolding under_def using 11 Refl + by (auto simp add: refl_on_def) + hence "b1 = b2" using BIJ * ** *** + by (simp add: bij_betw_def inj_on_def) } ultimately show "b1 = b2" - using Total by (auto simp add: total_on_def) + using Total by (auto simp add: total_on_def) qed - (* *) + (* *) let ?f' = "(inv_into A f)" - (* *) + (* *) have 2: "\b \ A. bij_betw ?f' (under r' (f b)) (under r b)" proof(clarify) fix b assume *: "b \ A" hence "under r b \ A" - using Well OF by(auto simp add: wo_rel.ofilter_def) + using Well OF by(auto simp add: wo_rel.ofilter_def) moreover have "f ` (under r b) = under r' (f b)" - using * BIJ by (auto simp add: bij_betw_def) + using * BIJ by (auto simp add: bij_betw_def) ultimately show "bij_betw ?f' (under r' (f b)) (under r b)" - using 1 by (auto simp add: bij_betw_inv_into_subset) + using 1 by (auto simp add: bij_betw_inv_into_subset) qed - (* *) + (* *) have 3: "\b' \ Field r'. bij_betw ?f' (under r' b') (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) + 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 + 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' (under r' b') (under r (?f' b'))" - using 2 by auto + using 2 by auto qed - (* *) + (* *) thus ?thesis unfolding embed_def . qed lemma inv_into_underS_embed: -assumes WELL: "Well_order r" and - BIJ: "\b \ underS r a. bij_betw f (under r b) (under r' (f b))" and - IN: "a \ Field r" and - IMAGE: "f ` (underS r a) = Field r'" -shows "embed r' r (inv_into (underS r a) f)" -using assms -by(auto simp add: wo_rel_def wo_rel.underS_ofilter inv_into_ofilter_embed) + assumes WELL: "Well_order r" and + BIJ: "\b \ underS r a. bij_betw f (under r b) (under r' (f b))" and + IN: "a \ Field r" and + IMAGE: "f ` (underS r a) = Field r'" + shows "embed r' r (inv_into (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)" + 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 (under r b) (under r' (f b)))" - using EMB by (auto simp add: embed_def) + using EMB by (auto simp add: embed_def) moreover have "f ` (Field r) \ Field r'" - using EMB WELL by (auto simp add: embed_Field) + 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) + by(auto simp add: wo_rel_def wo_rel.Field_ofilter inv_into_ofilter_embed) qed lemma inv_into_Field_embed_bij_betw: @@ -487,12 +487,12 @@ \ 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 \ underS r a. bij_betw f (under r b) (under r' (f b))" and - NOT: "f ` (underS r a) \ Field r'" and SUC: "f a = wo_rel.suc r' (f`(underS r a))" -shows "bij_betw f (under r a) (under r' (f a))" + 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 \ underS r a. bij_betw f (under r b) (under r' (f b))" and + NOT: "f ` (underS r a) \ Field r'" and SUC: "f a = wo_rel.suc r' (f`(underS r a))" + shows "bij_betw f (under r a) (under r' (f a))" proof- (* Preliminary facts *) have Well: "wo_rel r" using WELL unfolding wo_rel_def . @@ -500,197 +500,197 @@ 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 (underS r a)" - by (auto simp add: Well wo_rel.underS_ofilter) + by (auto simp add: Well wo_rel.underS_ofilter) hence UN: "underS r a = (\b \ underS r a. under r b)" - using Well wo_rel.ofilter_under_UNION[of r "underS r a"] by blast - (* Gather facts about elements of underS r a *) + using Well wo_rel.ofilter_under_UNION[of r "underS r a"] by blast + (* Gather facts about elements of underS r a *) {fix b assume *: "b \ underS r a" - hence t0: "(b,a) \ r \ b \ a" unfolding underS_def by auto - have t1: "b \ Field r" - using * underS_Field[of r a] by auto - have t2: "f`(under r b) = under r' (f b)" - using IH * by (auto simp add: bij_betw_def) - hence t3: "wo_rel.ofilter r' (f`(under r b))" - using Well' by (auto simp add: wo_rel.under_ofilter) - have "f`(under r b) \ Field r'" - using t2 by (auto simp add: under_Field) - moreover - have "b \ under r b" - using t1 by(auto simp add: Refl Refl_under_in) - ultimately - have t4: "f b \ Field r'" by auto - have "f`(under r b) = under r' (f b) \ + hence t0: "(b,a) \ r \ b \ a" unfolding underS_def by auto + have t1: "b \ Field r" + using * underS_Field[of r a] by auto + have t2: "f`(under r b) = under r' (f b)" + using IH * by (auto simp add: bij_betw_def) + hence t3: "wo_rel.ofilter r' (f`(under r b))" + using Well' by (auto simp add: wo_rel.under_ofilter) + have "f`(under r b) \ Field r'" + using t2 by (auto simp add: under_Field) + moreover + have "b \ under r b" + using t1 by(auto simp add: Refl Refl_under_in) + ultimately + have t4: "f b \ Field r'" by auto + have "f`(under r b) = under r' (f b) \ wo_rel.ofilter r' (f`(under r b)) \ f b \ Field r'" - using t2 t3 t4 by auto + using t2 t3 t4 by auto } hence bFact: - "\b \ underS r a. f`(under r b) = under r' (f b) \ + "\b \ underS r a. f`(under r b) = under r' (f b) \ wo_rel.ofilter r' (f`(under r b)) \ f b \ Field r'" by blast - (* *) + (* *) have subField: "f`(underS r a) \ Field r'" - using bFact by blast - (* *) + using bFact by blast + (* *) have OF': "wo_rel.ofilter r' (f`(underS r a))" proof- have "f`(underS r a) = f`(\b \ underS r a. under r b)" - using UN by auto + using UN by auto also have "\ = (\b \ underS r a. f`(under r b))" by blast also have "\ = (\b \ underS r a. (under r' (f b)))" - using bFact by auto + using bFact by auto finally have "f`(underS r a) = (\b \ underS r a. (under r' (f b)))" . thus ?thesis - using Well' bFact - wo_rel.ofilter_UNION[of r' "underS r a" "\ b. under r' (f b)"] by fastforce + using Well' bFact + wo_rel.ofilter_UNION[of r' "underS r a" "\ b. under r' (f b)"] by fastforce qed - (* *) + (* *) have "f`(underS r a) \ AboveS r' (f`(underS r a)) = Field r'" - using Well' OF' by (auto simp add: wo_rel.ofilter_AboveS_Field) + using Well' OF' by (auto simp add: wo_rel.ofilter_AboveS_Field) hence NE: "AboveS r' (f`(underS r a)) \ {}" - using subField NOT by blast - (* Main proof *) + using subField NOT by blast + (* Main proof *) have INCL1: "f`(underS r a) \ underS r' (f a) " proof(auto) fix b assume *: "b \ 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`(underS r a)" "f b"] by force + using subField Well' SUC NE * + wo_rel.suc_greater[of r' "f`(underS r a)" "f b"] by force thus "f b \ underS r' (f a)" - unfolding underS_def by simp + unfolding underS_def by simp qed - (* *) + (* *) have INCL2: "underS r' (f a) \ f`(underS r a)" proof fix b' assume "b' \ underS r' (f a)" hence "b' \ f a \ (b', f a) \ r'" - unfolding underS_def by simp + unfolding underS_def by simp thus "b' \ f`(underS r a)" - using Well' SUC NE OF' - wo_rel.suc_ofilter_in[of r' "f ` underS r a" b'] by auto + using Well' SUC NE OF' + wo_rel.suc_ofilter_in[of r' "f ` underS r a" b'] by auto qed - (* *) + (* *) have INJ: "inj_on f (underS r a)" proof- have "\b \ underS r a. inj_on f (under r b)" - using IH by (auto simp add: bij_betw_def) + using IH by (auto simp add: bij_betw_def) moreover have "\b. wo_rel.ofilter r (under r b)" - using Well by (auto simp add: wo_rel.under_ofilter) + using Well by (auto simp add: wo_rel.under_ofilter) ultimately show ?thesis - using WELL bFact UN - UNION_inj_on_ofilter[of r "underS r a" "\b. under r b" f] - by auto + using WELL bFact UN + UNION_inj_on_ofilter[of r "underS r a" "\b. under r b" f] + by auto qed - (* *) + (* *) have BIJ: "bij_betw f (underS r a) (underS r' (f a))" - unfolding bij_betw_def - using INJ INCL1 INCL2 by auto - (* *) + 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) + 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 + 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`(underS r a) \ f`(underS r a) \ Field r' + 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`(underS r a) \ f`(underS r a) \ Field r' \ f a = wo_rel.suc r' (f`(underS r a)) \ g a = True) \ (\(False \ (g`(underS r a)) \ f`(underS r a) \ Field r') \ g a = False)" and -MAIN2: "\ a. a \ Field r \ False \ g`(under r a) \ + MAIN2: "\ a. a \ Field r \ False \ g`(under r a) \ bij_betw f (under r a) (under r' (f a))" and -Case: "a \ Field r \ False \ g`(under r a)" -shows "\f'. embed r' r f'" + Case: "a \ Field r \ False \ g`(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: "under r a = underS r a \ {a}" - using Refl Case by(auto simp add: Refl_under_underS) - (* *) + using Refl Case by(auto simp add: Refl_under_underS) + (* *) have 1: "g a = False" proof- {assume "g a \ False" - with 0 Case have "False \ g`(underS r a)" by blast - with MAIN1 have "g a = False" by blast} + with 0 Case have "False \ g`(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`(underS r ?a)" proof assume "False \ g`(underS r ?a)" then obtain b where "b \ underS r ?a" and 31: "g b = False" by auto hence 32: "(b,?a) \ r \ b \ ?a" - by (auto simp add: underS_def) + by (auto simp add: 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? *) + (* again: why worked without type annotations? *) with 32 Antisym show False - by (auto simp add: antisym_def) + by (auto simp add: antisym_def) qed have temp: "?a \ ?A" - using Well 2 wo_rel.minim_in[of r ?A] by auto + 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`(underS r ?a) = Field r'" - using MAIN1[of ?a] 3 5 by blast - (* *) + using MAIN1[of ?a] 3 5 by blast + (* *) have 7: "\b \ underS r ?a. bij_betw f (under r b) (under r' (f b))" proof fix b assume as: "b \ underS r ?a" moreover have "wo_rel.ofilter r (underS r ?a)" - using Well by (auto simp add: wo_rel.underS_ofilter) + using Well by (auto simp add: wo_rel.underS_ofilter) ultimately have "False \ g`(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: underS_def) + unfolding Field_def using as by (auto simp add: underS_def) ultimately show "bij_betw f (under r b) (under r' (f b))" - using MAIN2 by auto + using MAIN2 by auto qed - (* *) + (* *) have "embed r' r (inv_into (underS r ?a) f)" - using WELL WELL' 7 4 6 inv_into_underS_embed[of r ?a f r'] by auto + using WELL WELL' 7 4 6 inv_into_underS_embed[of r ?a f r'] by auto thus ?thesis - unfolding embed_def by blast + 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')" + 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 *) + (* Main proof *) obtain H where H_def: "H = (\h a. if False \ (snd \ h)`(underS r a) \ (fst \ h)`(underS r a) \ Field r' then (wo_rel.suc r' ((fst \ h)`(underS r a)), True) else (undefined, False))" by blast have Adm: "wo_rel.adm_wo r H" - using Well + using Well proof(unfold wo_rel.adm_wo_def, clarify) fix h1::"'a \ 'a' * bool" and h2::"'a \ 'a' * bool" and x assume "\y\underS r x. h1 y = h2 y" @@ -701,52 +701,52 @@ 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: *) + (* 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 \ h" and g_def: "g = snd \ h" by blast + where h_def: "h = wo_rel.worec r H" and + f_def: "f = fst \ h" and g_def: "g = snd \ h" by blast obtain test where test_def: - "test = (\ a. False \ (g`(underS r a)) \ f`(underS r a) \ Field r')" by blast - (* *) + "test = (\ a. False \ (g`(underS r a)) \ f`(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) + 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`(underS r a)) \ g a = True) \ + "\ a. (test a \ f a = wo_rel.suc r' (f`(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`(underS r a)) \ g a = True) \ (\(test a) \ g a = False)" - using *[of a] test_def f_def g_def H_def by auto + using *[of a] test_def f_def g_def H_def by auto qed - (* *) + (* *) let ?phi = "\ a. a \ Field r \ False \ g`(under r a) \ bij_betw f (under r a) (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) + simp only: Well, clarify) fix a assume IH: "\b. b \ a \ (b,a) \ r \ ?phi b" and - *: "a \ Field r" and - **: "False \ g`(under r a)" + *: "a \ Field r" and + **: "False \ g`(under r a)" have 1: "\b \ underS r a. bij_betw f (under r b) (under r' (f b))" proof(clarify) fix b assume ***: "b \ underS r a" hence 0: "(b,a) \ r \ b \ a" unfolding underS_def by auto moreover have "b \ Field r" - using *** underS_Field[of r a] by auto + using *** underS_Field[of r a] by auto moreover have "False \ g`(under r b)" - using 0 ** Trans under_incr[of r b a] by auto + using 0 ** Trans under_incr[of r b a] by auto ultimately show "bij_betw f (under r b) (under r' (f b))" - using IH by auto + using IH by auto qed - (* *) + (* *) have 21: "False \ g`(underS r a)" - using ** underS_subset_under[of r a] by auto + using ** underS_subset_under[of r a] by auto have 22: "g`(under r a) \ {True}" using ** by auto moreover have 23: "a \ under r a" - using Refl * by (auto simp add: Refl_under_in) + using Refl * by (auto simp add: Refl_under_in) ultimately have 24: "g a = True" by blast have 2: "f`(underS r a) \ Field r'" proof @@ -754,29 +754,29 @@ 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`(underS r a))" - using 21 2 Main1 test_def by blast - (* *) + using 21 2 Main1 test_def by blast + (* *) show "bij_betw f (under r a) (under r' (f a))" - using WELL WELL' 1 2 3 * - wellorders_totally_ordered_aux[of r r' a f] by auto + 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`(under r a))" show ?thesis proof(cases "\a. ?chi a") assume "\ (\a. ?chi a)" hence "\a \ Field r. bij_betw f (under r a) (under r' (f a))" - using Main2 by blast + 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 fast + using wellorders_totally_ordered_aux2[of r r' g f a] + WELL WELL' Main1 Main2 test_def by fast thus ?thesis by blast qed qed @@ -790,75 +790,75 @@ 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`(underS r a))" + 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`(underS r a))" proof- have "bij_betw f (underS r a) (underS r' (f a))" - using assms by (auto simp add: embed_underS) + using assms by (auto simp add: embed_underS) hence "f`(underS r a) = underS r' (f a)" - by (auto simp add: bij_betw_def) + 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' (underS r' (f a))" - using WELL' by (auto simp add: wo_rel_def wo_rel.suc_underS) + using EMB WELL embed_Field[of r r' f] by auto + hence "f a = wo_rel.suc r' (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" + 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" + *: "a \ Field r" hence "\b \ underS r a. f b = g b" - unfolding underS_def by (auto simp add: Field_def) + unfolding underS_def by (auto simp add: Field_def) hence "f`(underS r a) = g`(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 + 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')" + 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' \ f)" using assms - by(auto simp add: comp_embed) + by(auto simp add: comp_embed) moreover have "embed r r id" using assms - by (auto simp add: id_embed) + by (auto simp add: id_embed) ultimately have "\a \ Field r. f'(f a) = a" - using assms embed_unique[of r r "f' \ f" id] id_def by auto + using assms embed_unique[of r r "f' \ f" id] id_def by auto moreover {have "embed r' r' (f \ 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 \ f'" id] id_def by auto + 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 \ 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')" + 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 \ f) \ embed r' r' (f \ g)" - using assms by (auto simp add: comp_embed) + 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 \ f" id] - WELL' id_embed[of r'] embed_unique[of r' r' "f \ g" id] - id_def by auto + using WELL id_embed[of r] embed_unique[of r r "g \ f" id] + WELL' id_embed[of r'] embed_unique[of r' r' "f \ 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 - (* *) + 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" @@ -872,24 +872,24 @@ 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) + 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')" + 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) + 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) + 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 @@ -928,7 +928,7 @@ hence 2: "embed r r'' ?g" using WELL EMB comp_embed[of r r' f r'' f'] by auto moreover have \
: "f' ` Field r' \ Field r''" - by (simp add: "1" embed_Field) + by (simp add: "1" embed_Field) {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) @@ -974,57 +974,57 @@ using assms unfolding iso_def 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'" + 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) + 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) + 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')" + 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) + 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 + unfolding bij_betw_def by blast thus "embedS r r' f" unfolding embedS_def - using ISO by auto + using ISO by auto qed lemma iso_Field: "iso r r' f \ f ` (Field r) = Field r'" 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')" + 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) + 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 + unfolding bij_betw_def by simp with * show "iso r r' f" unfolding iso_def by auto qed lemma iso_iff2: "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')" - (is "?lhs = ?rhs") + (is "?lhs = ?rhs") proof assume L: ?lhs then have "bij_betw f (Field r) (Field r')" and emb: "embed r r' f" @@ -1047,31 +1047,31 @@ 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)" + 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) + 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) + 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 + 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'" + ***: "(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 + 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 + using Well ** wo_rel.TOTAL[of r] total_on_def[of _ r] by blast qed qed @@ -1098,12 +1098,7 @@ lemma iso_forward: assumes "(x,y) \ r" "iso r r' f" shows "(f x,f y) \ r'" -proof - - have "x \ Field r" "y \ Field r" - using assms by (auto simp: Field_def) - with assms show ?thesis unfolding iso_iff2 by blast -qed - + using assms by (auto simp: Field_def iso_iff2) lemma iso_trans: assumes "trans r" and iso: "iso r r' f" shows "trans r'" @@ -1112,7 +1107,7 @@ fix x y z assume xyz: "(x, y) \ r'" "(y, z) \ r'" then have *: "(inv_into (Field r) f x, inv_into (Field r) f y) \ r" - "(inv_into (Field r) f y, inv_into (Field r) f z) \ r" + "(inv_into (Field r) f y, inv_into (Field r) f z) \ r" using iso_backward [OF _ iso] by blast+ then have "inv_into (Field r) f x \ Field r" "inv_into (Field r) f y \ Field r" by (auto simp: Field_def) diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/BNF_Wellorder_Relation.thy --- a/src/HOL/BNF_Wellorder_Relation.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/BNF_Wellorder_Relation.thy Sun Jan 15 20:00:44 2023 +0100 @@ -8,7 +8,7 @@ section \Well-Order Relations as Needed by Bounded Natural Functors\ theory BNF_Wellorder_Relation -imports Order_Relation + imports Order_Relation begin text\In this section, we develop basic concepts and results pertaining @@ -41,35 +41,35 @@ subsection \Auxiliaries\ lemma REFL: "Refl r" -using WELL order_on_defs[of _ r] by auto + using WELL order_on_defs[of _ r] by auto lemma TRANS: "trans r" -using WELL order_on_defs[of _ r] by auto + using WELL order_on_defs[of _ r] by auto lemma ANTISYM: "antisym r" -using WELL order_on_defs[of _ r] by auto + using WELL order_on_defs[of _ r] by auto lemma TOTAL: "Total r" -using WELL order_on_defs[of _ r] by auto + 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 + 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 + 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 + 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. \{a,b} <= Field r; ((a,b) \ r \ phi a b); ((b,a) \ r \ phi a b)\ \ phi a b" -using TOTALS by auto + 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); + "\ 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 + using TOTALS by auto subsection \Well-founded induction and recursion adapted to non-strict well-order relations\ @@ -81,34 +81,34 @@ 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" + 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 + 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" + 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" + 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)" + 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[of r] by auto + unfolding adm_wf_def + using ADM adm_wo_def[of H] underS_def[of r] by auto hence "wfrec ?rS H = H (wfrec ?rS H)" - using WF wfrec_fixpoint[of ?rS H] by simp + using WF wfrec_fixpoint[of ?rS H] by simp thus ?thesis unfolding worec_def . qed @@ -127,97 +127,83 @@ 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" + 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)" + where "isMinim A b \ b \ A \ (\a \ A. (b,a) \ r)" definition minim :: "'a set \ 'a" -where "minim A \ THE b. isMinim A b" + where "minim A \ THE b. isMinim A b" definition supr :: "'a set \ 'a" -where "supr A \ minim (Above A)" + where "supr A \ minim (Above A)" definition suc :: "'a set \ 'a" -where "suc A \ minim (AboveS A)" + where "suc A \ minim (AboveS 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}" + 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) + 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) + 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) + 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 + 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) + 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 + 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) + 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 + 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 lemma in_notinI: -assumes "(j,i) \ r \ j = i" and "i \ Field r" and "j \ Field r" -shows "(i,j) \ r" using assms max2_def max2_greater_among by fastforce + assumes "(j,i) \ r \ j = i" and "i \ Field r" and "j \ Field r" + shows "(i,j) \ r" using assms max2_def max2_greater_among by fastforce 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 + assumes "isMinim B a" "isMinim B a'" + shows "a = a'" + using assms ANTISYM antisym_def[of r] by (auto simp: isMinim_def) lemma Well_order_isMinim_exists: -assumes SUB: "B \ Field r" and NE: "B \ {}" -shows "\b. isMinim B b" + 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 @@ -232,52 +218,46 @@ moreover have "b' = b \ (b, b') \ r" using ** REFL by (auto simp add: refl_on_def) moreover have "b' \ b \ (b',b) \ r \ (b,b') \ r" - using ** TOTAL by (auto simp add: total_on_def) + using ** TOTAL by (auto simp add: total_on_def) ultimately show "(b,b') \ r" by blast qed qed - then have "isMinim B b" + then show ?thesis unfolding isMinim_def using * by auto - then show ?thesis - by auto qed lemma minim_isMinim: -assumes SUB: "B \ Field r" and NE: "B \ {}" -shows "isMinim B (minim B)" + 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 + using isMinim_unique * by auto ultimately show ?thesis - unfolding minim_def using theI[of ?phi b] by blast + 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 + assumes "B \ Field r" and "B \ {}" + shows "minim B \ B" + using assms minim_isMinim[of B] by (auto simp: isMinim_def) lemma minim_inField: -assumes "B \ Field r" and "B \ {}" -shows "minim B \ Field r" + 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" + 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 @@ -285,137 +265,104 @@ 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" + 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 + 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" + 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[of r] by auto + using AboveS_Field[of r] by auto thus "minim (AboveS B) \ AboveS B" - using assms by (simp add: minim_in) + 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[of r] show ?thesis by simp -qed + assumes SUB: "B \ Field r" and ABOVES: "AboveS B \ {}" and IN: "b \ B" + shows "suc B \ b \ (b,suc B) \ r" + using IN AboveS_def[of r] assms suc_AboveS by auto 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[of r] by auto - thus "(minim (AboveS B),a) \ r" - using assms minim_least by simp -qed + assumes ABOVES: "a \ AboveS B" + shows "(suc B,a) \ r" + using assms minim_least AboveS_Field[of r] by (auto simp: suc_def) 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[of r] by auto -qed + assumes "B \ Field r" and "AboveS B \ {}" + shows "suc B \ Field r" + using suc_AboveS assms AboveS_Field[of r] by auto 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 r B] by auto - thus "a = minim (AboveS B)" - using assms equals_minim - by simp -qed + assumes "B \ Field r" and "a \ AboveS B" and "\ a'. a' \ AboveS B \ (a,a') \ r" + shows "a = suc B" + using assms equals_minim AboveS_Field[of r B] by (auto simp: suc_def) lemma suc_underS: -assumes IN: "a \ Field r" -shows "a = suc (underS a)" + assumes IN: "a \ Field r" + shows "a = suc (underS a)" proof- have "underS a \ Field r" - using underS_Field[of r] by auto + using underS_Field[of r] by auto moreover have "a \ AboveS (underS a)" - using in_AboveS_underS IN by fast + using in_AboveS_underS IN by fast 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 fast + using AboveS_Field by fast {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 fast - with * have False by simp - } - ultimately have "(a,a') \ r" by blast + 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 fast + 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 + using equals_suc_AboveS by auto qed subsubsection \Properties of order filters\ -lemma under_ofilter: -"ofilter (under a)" -proof - - have "\aa x. (aa, a) \ r \ (x, aa) \ r \ (x, a) \ r" - proof - - fix aa x - assume "(aa,a) \ r" "(x,aa) \ r" - then show "(x,a) \ r" - using TRANS trans_def[of r] by blast - qed - then show ?thesis unfolding ofilter_def under_def - by (auto simp add: Field_def) -qed +lemma under_ofilter: "ofilter (under a)" + using TRANS by (auto simp: ofilter_def under_def Field_iff trans_def) -lemma underS_ofilter: -"ofilter (underS a)" +lemma underS_ofilter: "ofilter (underS a)" unfolding ofilter_def underS_def under_def proof safe - fix aa assume "(a, aa) \ r" "(aa, a) \ r" and DIFF: "aa \ a" + fix b assume "(a, b) \ r" "(b, a) \ r" and DIFF: "b \ a" thus False - using ANTISYM antisym_def[of r] by blast + using ANTISYM antisym_def[of r] by blast next - fix aa x - assume "(aa,a) \ r" "aa \ a" "(x,aa) \ r" + fix b x + assume "(b,a) \ r" "b \ a" "(x,b) \ r" thus "(x,a) \ r" using TRANS trans_def[of r] by blast next @@ -427,15 +374,15 @@ qed lemma Field_ofilter: -"ofilter (Field r)" -by(unfold ofilter_def under_def, auto simp add: Field_def) + "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))" + "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) + by (auto simp: underS_ofilter Field_ofilter) next assume *: "ofilter A" let ?One = "(\a\Field r. A = underS a)" @@ -451,7 +398,7 @@ 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 by auto - (* *) + (* *) moreover have "A = underS ?a" proof @@ -462,12 +409,12 @@ 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[of r] by auto - hence "?a \ A" using ** 13 by blast - with 4 have False by simp + 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[of r] by auto + hence "?a \ A" using ** 13 by blast + with 4 have False by simp } then have "(x,?a) \ r" by blast thus "x \ underS ?a" @@ -479,13 +426,13 @@ fix x assume **: "x \ underS ?a" hence 11: "x \ Field r" - using Field_def unfolding underS_def by fastforce - {assume "x \ A" + using Field_def unfolding underS_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] ** unfolding underS_def by auto - } + using ANTISYM antisym_def[of r] ** unfolding underS_def by auto + } thus "x \ A" by blast qed qed @@ -499,27 +446,27 @@ qed lemma ofilter_UNION: -"(\ i. i \ I \ ofilter(A i)) \ ofilter (\i \ I. A i)" -unfolding ofilter_def by blast + "(\ 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)" + assumes "ofilter A" + shows "A = (\a \ A. under a)" proof have "\a \ A. under a \ A" - using assms ofilter_def by auto + 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[of r] assms ofilter_def[of A] by blast + using REFL Refl_under_in[of r] assms ofilter_def[of A] 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" + 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 @@ -527,7 +474,7 @@ next assume Case2: "A \ Field r" with ofilter_underS_Field OF1 obtain a where - 1: "a \ Field r \ A = underS a" by auto + 1: "a \ Field r \ A = underS a" by auto show ?thesis proof(cases "B = Field r") assume Case21: "B = Field r" @@ -536,68 +483,68 @@ next assume Case22: "B \ Field r" with ofilter_underS_Field OF2 obtain b where - 2: "b \ Field r \ B = underS b" by auto + 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 + 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[of r] TRANS ANTISYM 1 2 - have "A \ B" by auto - hence ?thesis by auto + with underS_incr[of r] TRANS ANTISYM 1 2 + have "A \ B" by auto + hence ?thesis by auto } moreover - {assume "(b,a) \ r" - with underS_incr[of r] TRANS ANTISYM 1 2 - have "B \ A" by auto - hence ?thesis by auto + {assume "(b,a) \ r" + with underS_incr[of r] 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" + assumes "ofilter A" + shows "A \ (AboveS A) = Field r" proof show "A \ (AboveS A) \ Field r" - using assms ofilter_def AboveS_Field[of r] by auto + using assms ofilter_def AboveS_Field[of r] 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[of r] have "x \ A" by auto - with ** have False by contradiction + {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[of r] 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 } - 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 + 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" + 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[of "Field r"] by auto + using WELL REL well_order_on_domain[of "Field r"] 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) + 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 diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Basic_BNFs.thy --- a/src/HOL/Basic_BNFs.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Basic_BNFs.thy Sun Jan 15 20:00:44 2023 +0100 @@ -130,7 +130,7 @@ lemma rel_prod_conv: "rel_prod R1 R2 = (\(a, b) (c, d). R1 a c \ R2 b d)" - by (rule ext, rule ext) auto + by force definition pred_fun :: "('a \ bool) \ ('b \ bool) \ ('a \ 'b) \ bool" @@ -198,27 +198,22 @@ lemma regularCard_bd_fun: "regularCard (natLeq +c card_suc ( |UNIV| ))" (is "regularCard (_ +c card_suc ?U)") - apply (cases "Cinfinite ?U") - apply (rule regularCard_csum) - apply (rule natLeq_Cinfinite) - apply (rule Cinfinite_card_suc) - apply assumption - apply (rule card_of_card_order_on) - apply (rule regularCard_natLeq) - apply (rule regularCard_card_suc) - apply (rule card_of_card_order_on) - apply assumption - apply (rule regularCard_ordIso[of natLeq]) - apply (rule csum_absorb1[THEN ordIso_symmetric]) - apply (rule natLeq_Cinfinite) - apply (rule card_suc_least) - apply (rule card_of_card_order_on) - apply (rule natLeq_Card_order) - apply (subst finite_iff_ordLess_natLeq[symmetric]) - apply (simp add: cinfinite_def Field_card_of card_of_card_order_on) - apply (rule natLeq_Cinfinite) - apply (rule regularCard_natLeq) - done +proof (cases "Cinfinite ?U") + case True + then show ?thesis + by (intro regularCard_csum natLeq_Cinfinite Cinfinite_card_suc + card_of_card_order_on regularCard_natLeq regularCard_card_suc) +next + case False + then have "card_suc ?U \o natLeq" + unfolding cinfinite_def Field_card_of + by (intro card_suc_least; + simp add: natLeq_Card_order card_of_card_order_on flip: finite_iff_ordLess_natLeq) + then have "natLeq =o natLeq +c card_suc ?U" + using natLeq_Cinfinite csum_absorb1 ordIso_symmetric by blast + then show ?thesis + by (intro regularCard_ordIso[OF _ natLeq_Cinfinite regularCard_natLeq]) +qed lemma ordLess_bd_fun: "|UNIV::'a set| Cardinal Arithmetic\ theory Cardinal_Arithmetic -imports Cardinal_Order_Relation + imports Cardinal_Order_Relation begin subsection \Binary sum\ lemma csum_Cnotzero2: "Cnotzero r2 \ 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) + unfolding csum_def + by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE) lemma single_cone: "|{x}| =o cone" @@ -27,10 +27,10 @@ qed lemma cone_Cnotzero: "Cnotzero cone" -by (simp add: cone_not_czero Card_order_cone) + by (simp add: cone_not_czero Card_order_cone) lemma cone_ordLeq_ctwo: "cone \o ctwo" -unfolding cone_def ctwo_def card_of_ordLeq[symmetric] by auto + unfolding cone_def ctwo_def card_of_ordLeq[symmetric] by auto lemma csum_czero1: "Card_order r \ r +c czero =o r" unfolding czero_def csum_def Field_card_of @@ -44,7 +44,7 @@ subsection \Product\ lemma Times_cprod: "|A \ B| =o |A| *c |B|" -by (simp only: cprod_def Field_card_of card_of_refl) + by (simp only: cprod_def Field_card_of card_of_refl) lemma card_of_Times_singleton: fixes A :: "'a set" @@ -64,21 +64,20 @@ lemma cprod_cone: "Card_order r \ r *c cone =o r" unfolding cprod_def cone_def Field_card_of - by (drule card_of_Field_ordIso) (erule ordIso_transitive[OF card_of_Times_singleton]) - + by (metis (no_types) card_of_Field_ordIso card_of_Times_singleton ordIso_transitive) lemma ordLeq_cprod1: "\Card_order p1; Cnotzero p2\ \ p1 \o p1 *c p2" -unfolding cprod_def by (metis Card_order_Times1 czeroI) + unfolding cprod_def by (metis Card_order_Times1 czeroI) subsection \Exponentiation\ lemma cexp_czero: "r ^c czero =o cone" -unfolding cexp_def czero_def Field_card_of Func_empty by (rule single_cone) + unfolding cexp_def czero_def Field_card_of Func_empty by (rule single_cone) lemma Pow_cexp_ctwo: "|Pow A| =o ctwo ^c |A|" -unfolding ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func) + by (simp add: card_of_Pow_Func cexp_def ctwo_def) lemma Cnotzero_cexp: assumes "Cnotzero q" @@ -92,29 +91,24 @@ lemma Cinfinite_ctwo_cexp: "Cinfinite r \ Cinfinite (ctwo ^c r)" -unfolding ctwo_def cexp_def cinfinite_def Field_card_of -by (rule conjI, rule infinite_Func, auto) + unfolding ctwo_def cexp_def cinfinite_def Field_card_of + by (rule conjI, rule infinite_Func, auto) lemma cone_ordLeq_iff_Field: assumes "cone \o r" shows "Field r \ {}" -proof (rule ccontr) - assume "\ Field r \ {}" - hence "Field r = {}" by simp - thus False using card_of_empty3 - card_of_mono2[OF assms] Cnotzero_imp_not_empty[OF cone_Cnotzero] by auto -qed + by (metis assms card_of_empty3 card_of_mono2 cone_Cnotzero czeroI) lemma cone_ordLeq_cexp: "cone \o r1 \ cone \o r1 ^c r2" -by (simp add: cexp_def cone_def Func_non_emp cone_ordLeq_iff_Field) + by (simp add: cexp_def cone_def Func_non_emp cone_ordLeq_iff_Field) lemma Card_order_czero: "Card_order czero" -by (simp only: card_of_Card_order czero_def) + by (simp only: card_of_Card_order czero_def) lemma cexp_mono2'': assumes 2: "p2 \o r2" - and n1: "Cnotzero q" - and n2: "Card_order p2" + and n1: "Cnotzero q" + and n2: "Card_order p2" shows "q ^c p2 \o q ^c r2" proof (cases "p2 =o (czero :: 'a rel)") case True @@ -129,28 +123,20 @@ lemma csum_cexp: "\Cinfinite r1; Cinfinite r2; Card_order q; ctwo \o q\ \ q ^c r1 +c q ^c r2 \o q ^c (r1 +c r2)" apply (rule csum_cinfinite_bound) - apply (metis cexp_mono2' cinfinite_def finite.emptyI ordLeq_csum1) - apply (metis cexp_mono2' cinfinite_def finite.emptyI ordLeq_csum2) + apply (metis cexp_mono2' cinfinite_def finite.emptyI ordLeq_csum1) + apply (metis cexp_mono2' cinfinite_def finite.emptyI ordLeq_csum2) by (simp_all add: Card_order_cexp Cinfinite_csum1 Cinfinite_cexp cinfinite_cexp) lemma csum_cexp': "\Cinfinite r; Card_order q; ctwo \o q\ \ q +c r \o q ^c r" -apply (rule csum_cinfinite_bound) - apply (metis Cinfinite_Cnotzero ordLeq_cexp1) - apply (metis ordLeq_cexp2) - apply blast+ -by (metis Cinfinite_cexp) + apply (rule csum_cinfinite_bound) + apply (metis Cinfinite_Cnotzero ordLeq_cexp1) + apply (metis ordLeq_cexp2) + apply blast+ + by (metis Cinfinite_cexp) lemma card_of_Sigma_ordLeq_Cinfinite: "\Cinfinite r; |I| \o r; \i \ I. |A i| \o r\ \ |SIGMA i : I. A i| \o r" -unfolding cinfinite_def by (blast intro: card_of_Sigma_ordLeq_infinite_Field) - -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 -qed + unfolding cinfinite_def by (blast intro: card_of_Sigma_ordLeq_infinite_Field) lemma Cinfinite_ordLess_cexp: assumes r: "Cinfinite r" @@ -165,19 +151,19 @@ lemma infinite_ordLeq_cexp: assumes "Cinfinite r" shows "r \o r ^c r" -by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]]) + by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]]) lemma czero_cexp: "Cnotzero r \ czero ^c r =o czero" - by (drule Cnotzero_imp_not_empty) (simp add: cexp_def czero_def card_of_empty_ordIso) + by (metis Cnotzero_imp_not_empty cexp_def czero_def card_of_empty_ordIso Field_card_of Func_is_emp) lemma Func_singleton: -fixes x :: 'b and A :: "'a set" -shows "|Func A {x}| =o |{x}|" + fixes x :: 'b and A :: "'a set" + shows "|Func A {x}| =o |{x}|" proof (rule ordIso_symmetric) define f where [abs_def]: "f y a = (if y = x \ a \ A then x else undefined)" for y a have "Func A {x} \ f ` {x}" unfolding f_def Func_def by (force simp: fun_eq_iff) - hence "bij_betw f {x} (Func A {x})" unfolding bij_betw_def inj_on_def f_def Func_def - by (auto split: if_split_asm) + hence "bij_betw f {x} (Func A {x})" + unfolding bij_betw_def inj_on_def f_def Func_def by (auto split: if_split_asm) thus "|{x}| =o |Func A {x}|" using card_of_ordIso by blast qed @@ -225,32 +211,32 @@ 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) + 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) + 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) + unfolding cpow_def cinfinite_def by simp lemma Card_order_cpow: "Card_order (cpow r)" -unfolding cpow_def by (rule card_of_Card_order) + unfolding cpow_def by (rule card_of_Card_order) 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) + unfolding cpow_def by (metis Card_order_Pow cardSuc_ordLess_ordLeq card_of_Card_order) lemma cpow_cexp_ctwo: "cpow r =o ctwo ^c r" -unfolding cpow_def ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func) + unfolding cpow_def ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func) subsection \Inverse image\ lemma vimage_ordLeq: -assumes "|A| \o k" and "\ a \ A. |vimage f {a}| \o k" and "Cinfinite k" -shows "|vimage f A| \o k" + assumes "|A| \o k" and "\ a \ A. |vimage f {a}| \o k" and "Cinfinite k" + shows "|vimage f A| \o k" proof- have "vimage f A = (\a \ A. vimage f {a})" by auto also have "|\a \ A. vimage f {a}| \o k" - using UNION_Cinfinite_bound[OF assms] . + using UNION_Cinfinite_bound[OF assms] . finally show ?thesis . qed @@ -268,7 +254,8 @@ lemma cmax1: assumes "Card_order r" "Card_order s" "s \o r" shows "cmax r s =o r" -unfolding cmax_def proof (split if_splits, intro conjI impI) + unfolding cmax_def +proof (split if_splits, intro conjI impI) assume "cinfinite r \ cinfinite s" hence Cinf: "Cinfinite r" using assms(1,3) by (metis cinfinite_mono) have "czero +c r +c s =o r +c s" by (rule csum_czero2[OF Card_order_csum]) @@ -299,33 +286,14 @@ shows "cmax r s =o s" by (metis assms cmax1 cmax_com ordIso_transitive) -lemma csum_absorb2: "Cinfinite r2 \ r1 \o r2 \ r1 +c r2 =o r2" - by (metis csum_absorb2') - -lemma cprod_infinite2': "\Cnotzero r1; Cinfinite r2; r1 \o r2\ \ r1 *c r2 =o r2" - unfolding ordIso_iff_ordLeq - by (intro conjI cprod_cinfinite_bound ordLeq_cprod2 ordLeq_refl) - (auto dest!: ordIso_imp_ordLeq not_ordLeq_ordLess simp: czero_def Card_order_empty) - context fixes r s assumes r: "Cinfinite r" - and s: "Cinfinite s" + and s: "Cinfinite s" begin lemma cmax_csum: "cmax r s =o r +c s" -proof (cases "r \o s") - case True - hence "cmax r s =o s" by (metis cmax2 r s) - also have "s =o r +c s" by (metis True csum_absorb2 ordIso_symmetric s) - finally show ?thesis . -next - case False - hence "s \o r" by (metis ordLeq_total r s card_order_on_def) - hence "cmax r s =o r" by (metis cmax1 r s) - also have "r =o r +c s" by (metis \s \o r\ csum_absorb1 ordIso_symmetric r) - finally show ?thesis . -qed + by (simp add: Card_order_csum cmax_def csum_czero2 r) lemma cmax_cprod: "cmax r s =o r *c s" proof (cases "r \o s") @@ -344,39 +312,21 @@ end lemma Card_order_cmax: -assumes r: "Card_order r" and s: "Card_order s" -shows "Card_order (cmax r s)" -unfolding cmax_def by (auto simp: Card_order_csum) + assumes r: "Card_order r" and s: "Card_order s" + shows "Card_order (cmax r s)" + unfolding cmax_def by (auto simp: Card_order_csum) lemma ordLeq_cmax: -assumes r: "Card_order r" and s: "Card_order s" -shows "r \o cmax r s \ s \o cmax r s" -proof- - {assume "r \o s" - hence ?thesis by (metis cmax2 ordIso_iff_ordLeq ordLeq_transitive r s) - } - moreover - {assume "s \o r" - hence ?thesis using cmax_com by (metis cmax2 ordIso_iff_ordLeq ordLeq_transitive r s) - } - ultimately show ?thesis using r s ordLeq_total unfolding card_order_on_def by auto -qed + assumes r: "Card_order r" and s: "Card_order s" + shows "r \o cmax r s \ s \o cmax r s" + by (meson card_order_on_def cmax1 cmax2 ordIso_iff_ordLeq ordLeq_total ordLeq_transitive r s) lemmas ordLeq_cmax1 = ordLeq_cmax[THEN conjunct1] and - ordLeq_cmax2 = ordLeq_cmax[THEN conjunct2] + ordLeq_cmax2 = ordLeq_cmax[THEN conjunct2] lemma finite_cmax: -assumes r: "Card_order r" and s: "Card_order s" -shows "finite (Field (cmax r s)) \ finite (Field r) \ finite (Field s)" -proof- - {assume "r \o s" - hence ?thesis by (metis cmax2 ordIso_finite_Field ordLeq_finite_Field r s) - } - moreover - {assume "s \o r" - hence ?thesis by (metis cmax1 ordIso_finite_Field ordLeq_finite_Field r s) - } - ultimately show ?thesis using r s ordLeq_total unfolding card_order_on_def by auto -qed + assumes r: "Card_order r" and s: "Card_order s" + shows "finite (Field (cmax r s)) \ finite (Field r) \ finite (Field s)" + by (meson card_order_on_def cmax1 cmax2 ordIso_finite_Field ordLeq_finite_Field ordLeq_total r s) end diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Cardinals/Cardinal_Order_Relation.thy --- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Sun Jan 15 20:00:44 2023 +0100 @@ -8,7 +8,7 @@ section \Cardinal-Order Relations\ theory Cardinal_Order_Relation -imports Wellorder_Constructions + imports Wellorder_Constructions begin declare @@ -23,7 +23,6 @@ card_of_mono1[simp] card_of_mono2[simp] card_of_cong[simp] - card_of_Field_ordLess[simp] card_of_Field_ordIso[simp] card_of_underS[simp] ordLess_Field[simp] @@ -72,46 +71,36 @@ subsection \Cardinal of a set\ lemma card_of_inj_rel: assumes INJ: "\x y y'. \(x,y) \ R; (x,y') \ R\ \ y = y'" -shows "|{y. \x. (x,y) \ R}| <=o |{x. \y. (x,y) \ R}|" + shows "|{y. \x. (x,y) \ R}| <=o |{x. \y. (x,y) \ R}|" proof- let ?Y = "{y. \x. (x,y) \ R}" let ?X = "{x. \y. (x,y) \ R}" let ?f = "\y. SOME x. (x,y) \ R" have "?f ` ?Y <= ?X" by (auto dest: someI) moreover have "inj_on ?f ?Y" - unfolding inj_on_def proof(auto) - fix y1 x1 y2 x2 - assume *: "(x1, y1) \ R" "(x2, y2) \ R" and **: "?f y1 = ?f y2" - hence "(?f y1,y1) \ R" using someI[of "\x. (x,y1) \ R"] by auto - moreover have "(?f y2,y2) \ R" using * someI[of "\x. (x,y2) \ R"] by auto - ultimately show "y1 = y2" using ** INJ by auto - qed + by (metis (mono_tags, lifting) assms inj_onI mem_Collect_eq) ultimately show "|?Y| <=o |?X|" using card_of_ordLeq by blast qed lemma card_of_unique2: "\card_order_on B r; bij_betw f A B\ \ r =o |A|" -using card_of_ordIso card_of_unique ordIso_equivalence by blast + using card_of_ordIso card_of_unique ordIso_equivalence by blast lemma internalize_card_of_ordLess2: -"( |A| B < C. |A| =o |B| \ |B| B < C. |A| =o |B| \ |B| {}" and "\r\R. Card_order r" -shows "Card_order (omax R)" -proof- - have "\r\R. Well_order r" - using assms unfolding card_order_on_def by simp - thus ?thesis using assms apply - apply(drule omax_in) by auto -qed + assumes "finite R" and "R \ {}" and "\r\R. Card_order r" + shows "Card_order (omax R)" + by (simp add: assms omax_in) lemma Card_order_omax2: -assumes "finite I" and "I \ {}" -shows "Card_order (omax {|A i| | i. i \ I})" + assumes "finite I" and "I \ {}" + shows "Card_order (omax {|A i| | i. i \ I})" proof- let ?R = "{|A i| | i. i \ I}" have "finite ?R" and "?R \ {}" using assms by auto moreover have "\r\?R. Card_order r" - using card_of_Card_order by auto + using card_of_Card_order by auto ultimately show ?thesis by(rule Card_order_omax) qed @@ -119,138 +108,104 @@ subsection \Cardinals versus set operations on arbitrary sets\ lemma card_of_set_type[simp]: "|UNIV::'a set| o |A \ B| " -using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce +lemma card_of_Un1[simp]: "|A| \o |A \ B| " + by fastforce -lemma card_of_diff[simp]: -shows "|A - B| \o |A|" -using inj_on_id[of "A - B"] card_of_ordLeq[of "A - B" _] by fastforce +lemma card_of_diff[simp]: "|A - B| \o |A|" + by fastforce lemma subset_ordLeq_strict: -assumes "A \ B" and "|A| (A < B)" - hence "A = B" using assms(1) by blast - hence False using assms(2) not_ordLess_ordIso card_of_refl by blast - } - thus ?thesis by blast -qed + assumes "A \ B" and "|A| B" and "|A| {}" -using assms subset_ordLeq_strict by blast + assumes "A \ B" and "|A| {}" + using assms subset_ordLeq_strict by blast lemma card_of_empty4: -"|{}::'b set| {})" -proof(intro iffI notI) - assume *: "|{}::'b set| {}" - hence "(\ (\f. inj_on f A \ f ` A \ {}))" - unfolding inj_on_def by blast - thus "| {} | {})" + by (metis card_of_empty card_of_ordLess2 image_is_empty not_ordLess_ordLeq) lemma card_of_empty5: -"|A| B \ {}" -using card_of_empty not_ordLess_ordLeq by blast + "|A| B \ {}" + using card_of_empty not_ordLess_ordLeq by blast lemma Well_order_card_of_empty: -"Well_order r \ |{}| \o r" by simp + "Well_order r \ |{}| \o r" + by simp lemma card_of_UNIV[simp]: -"|A :: 'a set| \o |UNIV :: 'a set|" -using card_of_mono1[of A] by simp + "|A :: 'a set| \o |UNIV :: 'a set|" + by simp lemma card_of_UNIV2[simp]: -"Card_order r \ (r :: 'a rel) \o |UNIV :: 'a set|" -using card_of_UNIV[of "Field r"] card_of_Field_ordIso - ordIso_symmetric ordIso_ordLeq_trans by blast + "Card_order r \ (r :: 'a rel) \o |UNIV :: 'a set|" + using card_of_UNIV[of "Field r"] card_of_Field_ordIso + ordIso_symmetric ordIso_ordLeq_trans by blast lemma card_of_Pow_mono[simp]: -assumes "|A| \o |B|" -shows "|Pow A| \o |Pow B|" + assumes "|A| \o |B|" + shows "|Pow A| \o |Pow B|" proof- obtain f where "inj_on f A \ f ` A \ B" - using assms card_of_ordLeq[of A B] by auto + using assms card_of_ordLeq[of A B] by auto hence "inj_on (image f) (Pow A) \ (image f) ` (Pow A) \ (Pow B)" - by (auto simp add: inj_on_image_Pow image_Pow_mono) + by (auto simp: inj_on_image_Pow image_Pow_mono) thus ?thesis using card_of_ordLeq[of "Pow A"] by metis qed lemma ordIso_Pow_mono[simp]: -assumes "r \o r'" -shows "|Pow(Field r)| \o |Pow(Field r')|" -using assms card_of_mono2 card_of_Pow_mono by blast + assumes "r \o r'" + shows "|Pow(Field r)| \o |Pow(Field r')|" + using assms card_of_mono2 card_of_Pow_mono by blast lemma card_of_Pow_cong[simp]: -assumes "|A| =o |B|" -shows "|Pow A| =o |Pow B|" -proof- - obtain f where "bij_betw f A B" - using assms card_of_ordIso[of A B] by auto - hence "bij_betw (image f) (Pow A) (Pow B)" - by (auto simp add: bij_betw_image_Pow) - thus ?thesis using card_of_ordIso[of "Pow A"] by auto -qed + assumes "|A| =o |B|" + shows "|Pow A| =o |Pow B|" + by (meson assms card_of_Pow_mono ordIso_iff_ordLeq) lemma ordIso_Pow_cong[simp]: -assumes "r =o r'" -shows "|Pow(Field r)| =o |Pow(Field r')|" -using assms card_of_cong card_of_Pow_cong by blast + assumes "r =o r'" + shows "|Pow(Field r)| =o |Pow(Field r')|" + using assms card_of_cong card_of_Pow_cong by blast corollary Card_order_Plus_empty1: -"Card_order r \ r =o |(Field r) <+> {}|" -using card_of_Plus_empty1 card_of_Field_ordIso ordIso_equivalence by blast + "Card_order r \ r =o |(Field r) <+> {}|" + using card_of_Plus_empty1 card_of_Field_ordIso ordIso_equivalence by blast corollary Card_order_Plus_empty2: -"Card_order r \ r =o |{} <+> (Field r)|" -using card_of_Plus_empty2 card_of_Field_ordIso ordIso_equivalence by blast - -lemma Card_order_Un1: -shows "Card_order r \ |Field r| \o |(Field r) \ B| " -using card_of_Un1 card_of_Field_ordIso ordIso_symmetric ordIso_ordLeq_trans by auto + "Card_order r \ r =o |{} <+> (Field r)|" + using card_of_Plus_empty2 card_of_Field_ordIso ordIso_equivalence by blast lemma card_of_Un2[simp]: -shows "|A| \o |B \ A|" -using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce - -lemma Card_order_Un2: -shows "Card_order r \ |Field r| \o |A \ (Field r)| " -using card_of_Un2 card_of_Field_ordIso ordIso_symmetric ordIso_ordLeq_trans by auto + shows "|A| \o |B \ A|" + by fastforce lemma Un_Plus_bij_betw: -assumes "A Int B = {}" -shows "\f. bij_betw f (A \ B) (A <+> B)" + assumes "A Int B = {}" + shows "\f. bij_betw f (A \ B) (A <+> B)" proof- - let ?f = "\ c. if c \ A then Inl c else Inr c" - have "bij_betw ?f (A \ B) (A <+> B)" - using assms by(unfold bij_betw_def inj_on_def, auto) + have "bij_betw (\ c. if c \ A then Inl c else Inr c) (A \ B) (A <+> B)" + using assms unfolding bij_betw_def inj_on_def by auto thus ?thesis by blast qed lemma card_of_Un_Plus_ordIso: -assumes "A Int B = {}" -shows "|A \ B| =o |A <+> B|" -using assms card_of_ordIso[of "A \ B"] Un_Plus_bij_betw[of A B] by auto + assumes "A Int B = {}" + shows "|A \ B| =o |A <+> B|" + by (meson Un_Plus_bij_betw assms card_of_ordIso) lemma card_of_Un_Plus_ordIso1: -"|A \ B| =o |A <+> (B - A)|" -using card_of_Un_Plus_ordIso[of A "B - A"] by auto + "|A \ B| =o |A <+> (B - A)|" + using card_of_Un_Plus_ordIso[of A "B - A"] by auto lemma card_of_Un_Plus_ordIso2: -"|A \ B| =o |(A - B) <+> B|" -using card_of_Un_Plus_ordIso[of "A - B" B] by auto + "|A \ B| =o |(A - B) <+> B|" + using card_of_Un_Plus_ordIso[of "A - B" B] by auto lemma card_of_Times_singl1: "|A| =o |A \ {b}|" proof- @@ -259,18 +214,19 @@ qed corollary Card_order_Times_singl1: -"Card_order r \ r =o |(Field r) \ {b}|" -using card_of_Times_singl1[of _ b] card_of_Field_ordIso ordIso_equivalence by blast + "Card_order r \ r =o |(Field r) \ {b}|" + using card_of_Times_singl1[of _ b] card_of_Field_ordIso ordIso_equivalence by blast lemma card_of_Times_singl2: "|A| =o |{b} \ A|" proof- - have "bij_betw snd ({b} \ A) A" unfolding bij_betw_def inj_on_def by force + have "bij_betw snd ({b} \ A) A" + unfolding bij_betw_def inj_on_def by force thus ?thesis using card_of_ordIso ordIso_symmetric by blast qed corollary Card_order_Times_singl2: -"Card_order r \ r =o |{a} \ (Field r)|" -using card_of_Times_singl2[of _ a] card_of_Field_ordIso ordIso_equivalence by blast + "Card_order r \ r =o |{a} \ (Field r)|" + using card_of_Times_singl2[of _ a] card_of_Field_ordIso ordIso_equivalence by blast lemma card_of_Times_assoc: "|(A \ B) \ C| =o |A \ B \ C|" proof - @@ -284,118 +240,111 @@ thus "x \ ?f ` ((A \ B) \ C)" by blast qed hence "bij_betw ?f ((A \ B) \ C) (A \ B \ C)" - unfolding bij_betw_def inj_on_def by auto + unfolding bij_betw_def inj_on_def by auto thus ?thesis using card_of_ordIso by blast qed -corollary Card_order_Times3: -"Card_order r \ |Field r| \o |(Field r) \ (Field r)|" - by (rule card_of_Times3) - lemma card_of_Times_cong1[simp]: -assumes "|A| =o |B|" -shows "|A \ C| =o |B \ C|" -using assms by (simp add: ordIso_iff_ordLeq) + assumes "|A| =o |B|" + shows "|A \ C| =o |B \ C|" + using assms by (simp add: ordIso_iff_ordLeq) lemma card_of_Times_cong2[simp]: -assumes "|A| =o |B|" -shows "|C \ A| =o |C \ B|" -using assms by (simp add: ordIso_iff_ordLeq) + assumes "|A| =o |B|" + shows "|C \ A| =o |C \ B|" + using assms by (simp add: ordIso_iff_ordLeq) lemma card_of_Times_mono[simp]: -assumes "|A| \o |B|" and "|C| \o |D|" -shows "|A \ C| \o |B \ D|" -using assms card_of_Times_mono1[of A B C] card_of_Times_mono2[of C D B] - ordLeq_transitive[of "|A \ C|"] by blast + assumes "|A| \o |B|" and "|C| \o |D|" + shows "|A \ C| \o |B \ D|" + using assms card_of_Times_mono1[of A B C] card_of_Times_mono2[of C D B] + ordLeq_transitive[of "|A \ C|"] by blast corollary ordLeq_Times_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_Times_mono by blast + 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_Times_mono by blast corollary ordIso_Times_cong1[simp]: -assumes "r =o r'" -shows "|(Field r) \ C| =o |(Field r') \ C|" -using assms card_of_cong card_of_Times_cong1 by blast + assumes "r =o r'" + shows "|(Field r) \ C| =o |(Field r') \ C|" + using assms card_of_cong card_of_Times_cong1 by blast corollary ordIso_Times_cong2: -assumes "r =o r'" -shows "|A \ (Field r)| =o |A \ (Field r')|" -using assms card_of_cong card_of_Times_cong2 by blast + assumes "r =o r'" + shows "|A \ (Field r)| =o |A \ (Field r')|" + using assms card_of_cong card_of_Times_cong2 by blast lemma card_of_Times_cong[simp]: -assumes "|A| =o |B|" and "|C| =o |D|" -shows "|A \ C| =o |B \ D|" -using assms -by (auto simp add: ordIso_iff_ordLeq) + assumes "|A| =o |B|" and "|C| =o |D|" + shows "|A \ C| =o |B \ D|" + using assms + by (auto simp: ordIso_iff_ordLeq) corollary ordIso_Times_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_Times_cong by blast + 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_Times_cong by blast lemma card_of_Sigma_mono2: -assumes "inj_on f (I::'i set)" and "f ` I \ (J::'j set)" -shows "|SIGMA i : I. (A::'j \ 'a set) (f i)| \o |SIGMA j : J. A j|" + assumes "inj_on f (I::'i set)" and "f ` I \ (J::'j set)" + shows "|SIGMA i : I. (A::'j \ 'a set) (f i)| \o |SIGMA j : J. A j|" proof- let ?LEFT = "SIGMA i : I. A (f i)" let ?RIGHT = "SIGMA j : J. A j" obtain u where u_def: "u = (\(i::'i,a::'a). (f i,a))" by blast have "inj_on u ?LEFT \ u `?LEFT \ ?RIGHT" - using assms unfolding u_def inj_on_def by auto + using assms unfolding u_def inj_on_def by auto thus ?thesis using card_of_ordLeq by blast qed lemma card_of_Sigma_mono: -assumes INJ: "inj_on f I" and IM: "f ` I \ J" and - LEQ: "\j \ J. |A j| \o |B j|" -shows "|SIGMA i : I. A (f i)| \o |SIGMA j : J. B j|" + assumes INJ: "inj_on f I" and IM: "f ` I \ J" and + LEQ: "\j \ J. |A j| \o |B j|" + shows "|SIGMA i : I. A (f i)| \o |SIGMA j : J. B j|" proof- have "\i \ I. |A(f i)| \o |B(f i)|" - using IM LEQ by blast + using IM LEQ by blast hence "|SIGMA i : I. A (f i)| \o |SIGMA i : I. B (f i)|" - using card_of_Sigma_mono1[of I] by metis + using card_of_Sigma_mono1[of I] by metis moreover have "|SIGMA i : I. B (f i)| \o |SIGMA j : J. B j|" - using INJ IM card_of_Sigma_mono2 by blast + using INJ IM card_of_Sigma_mono2 by blast ultimately show ?thesis using ordLeq_transitive by blast qed lemma ordLeq_Sigma_mono1: -assumes "\i \ I. p i \o r i" -shows "|SIGMA i : I. Field(p i)| \o |SIGMA i : I. Field(r i)|" -using assms by (auto simp add: card_of_Sigma_mono1) + assumes "\i \ I. p i \o r i" + shows "|SIGMA i : I. Field(p i)| \o |SIGMA i : I. Field(r i)|" + using assms by (auto simp: card_of_Sigma_mono1) lemma ordLeq_Sigma_mono: -assumes "inj_on f I" and "f ` I \ J" and - "\j \ J. p j \o r j" -shows "|SIGMA i : I. Field(p(f i))| \o |SIGMA j : J. Field(r j)|" -using assms card_of_mono2 card_of_Sigma_mono - [of f I J "\ i. Field(p i)" "\ j. Field(r j)"] by metis + assumes "inj_on f I" and "f ` I \ J" and + "\j \ J. p j \o r j" + shows "|SIGMA i : I. Field(p(f i))| \o |SIGMA j : J. Field(r j)|" + using assms card_of_mono2 card_of_Sigma_mono [of f I J "\ i. Field(p i)"] by metis lemma ordIso_Sigma_cong1: -assumes "\i \ I. p i =o r i" -shows "|SIGMA i : I. Field(p i)| =o |SIGMA i : I. Field(r i)|" -using assms by (auto simp add: card_of_Sigma_cong1) + assumes "\i \ I. p i =o r i" + shows "|SIGMA i : I. Field(p i)| =o |SIGMA i : I. Field(r i)|" + using assms by (auto simp: card_of_Sigma_cong1) lemma ordLeq_Sigma_cong: -assumes "bij_betw f I J" and - "\j \ J. p j =o r j" -shows "|SIGMA i : I. Field(p(f i))| =o |SIGMA j : J. Field(r j)|" -using assms card_of_cong card_of_Sigma_cong - [of f I J "\ j. Field(p j)" "\ j. Field(r j)"] by blast + assumes "bij_betw f I J" and + "\j \ J. p j =o r j" + shows "|SIGMA i : I. Field(p(f i))| =o |SIGMA j : J. Field(r j)|" + using assms card_of_cong card_of_Sigma_cong + [of f I J "\ j. Field(p j)" "\ j. Field(r j)"] by blast lemma card_of_UNION_Sigma2: -assumes -"!! i j. \{i,j} <= I; i ~= j\ \ A i Int A j = {}" -shows -"|\i\I. A i| =o |Sigma I A|" + assumes "\i j. \{i,j} <= I; i \ j\ \ A i Int A j = {}" + shows "|\i\I. A i| =o |Sigma I A|" proof- let ?L = "\i\I. A i" let ?R = "Sigma I A" have "|?L| <=o |?R|" using card_of_UNION_Sigma . moreover have "|?R| <=o |?L|" proof- have "inj_on snd ?R" - unfolding inj_on_def using assms by auto + unfolding inj_on_def using assms by auto moreover have "snd ` ?R <= ?L" by auto ultimately show ?thesis using card_of_ordLeq by blast qed @@ -403,102 +352,91 @@ qed corollary Plus_into_Times: -assumes A2: "a1 \ a2 \ {a1,a2} \ A" and - B2: "b1 \ b2 \ {b1,b2} \ B" -shows "\f. inj_on f (A <+> B) \ f ` (A <+> B) \ A \ B" -using assms by (auto simp add: card_of_Plus_Times card_of_ordLeq) + assumes A2: "a1 \ a2 \ {a1,a2} \ A" and B2: "b1 \ b2 \ {b1,b2} \ B" + shows "\f. inj_on f (A <+> B) \ f ` (A <+> B) \ A \ B" + using assms by (auto simp: card_of_Plus_Times card_of_ordLeq) corollary Plus_into_Times_types: -assumes A2: "(a1::'a) \ a2" and B2: "(b1::'b) \ b2" -shows "\(f::'a + 'b \ 'a * 'b). inj f" -using assms Plus_into_Times[of a1 a2 UNIV b1 b2 UNIV] -by auto + assumes A2: "(a1::'a) \ a2" and B2: "(b1::'b) \ b2" + shows "\(f::'a + 'b \ 'a * 'b). inj f" + using assms Plus_into_Times[of a1 a2 UNIV b1 b2 UNIV] + by auto corollary Times_same_infinite_bij_betw: -assumes "\finite A" -shows "\f. bij_betw f (A \ A) A" -using assms by (auto simp add: card_of_ordIso) + assumes "\finite A" + shows "\f. bij_betw f (A \ A) A" + using assms by (auto simp: card_of_ordIso) corollary Times_same_infinite_bij_betw_types: -assumes INF: "\finite(UNIV::'a set)" -shows "\(f::('a * 'a) => 'a). bij f" -using assms Times_same_infinite_bij_betw[of "UNIV::'a set"] -by auto + assumes INF: "\finite(UNIV::'a set)" + shows "\(f::('a * 'a) => 'a). bij f" + using assms Times_same_infinite_bij_betw[of "UNIV::'a set"] + by auto corollary Times_infinite_bij_betw: -assumes INF: "\finite A" and NE: "B \ {}" and INJ: "inj_on g B \ g ` B \ A" -shows "(\f. bij_betw f (A \ B) A) \ (\h. bij_betw h (B \ A) A)" + assumes INF: "\finite A" and NE: "B \ {}" and INJ: "inj_on g B \ g ` B \ A" + shows "(\f. bij_betw f (A \ B) A) \ (\h. bij_betw h (B \ A) A)" proof- have "|B| \o |A|" using INJ card_of_ordLeq by blast thus ?thesis using INF NE - by (auto simp add: card_of_ordIso card_of_Times_infinite) + by (auto simp: card_of_ordIso card_of_Times_infinite) qed corollary Times_infinite_bij_betw_types: -assumes INF: "\finite(UNIV::'a set)" and - BIJ: "inj(g::'b \ 'a)" -shows "(\(f::('b * 'a) => 'a). bij f) \ (\(h::('a * 'b) => 'a). bij h)" -using assms Times_infinite_bij_betw[of "UNIV::'a set" "UNIV::'b set" g] -by auto + assumes "\finite(UNIV::'a set)" and "inj(g::'b \ 'a)" + shows "(\(f::('b * 'a) => 'a). bij f) \ (\(h::('a * 'b) => 'a). bij h)" + using assms Times_infinite_bij_betw[of "UNIV::'a set" "UNIV::'b set" g] + by auto lemma card_of_Times_ordLeq_infinite: -"\\finite C; |A| \o |C|; |B| \o |C|\ - \ |A \ B| \o |C|" -by(simp add: card_of_Sigma_ordLeq_infinite) + "\\finite C; |A| \o |C|; |B| \o |C|\ \ |A \ B| \o |C|" + by(simp add: card_of_Sigma_ordLeq_infinite) corollary Plus_infinite_bij_betw: -assumes INF: "\finite A" and INJ: "inj_on g B \ g ` B \ A" -shows "(\f. bij_betw f (A <+> B) A) \ (\h. bij_betw h (B <+> A) A)" + assumes INF: "\finite A" and INJ: "inj_on g B \ g ` B \ A" + shows "(\f. bij_betw f (A <+> B) A) \ (\h. bij_betw h (B <+> A) A)" proof- have "|B| \o |A|" using INJ card_of_ordLeq by blast thus ?thesis using INF - by (auto simp add: card_of_ordIso) + by (auto simp: card_of_ordIso) qed corollary Plus_infinite_bij_betw_types: -assumes INF: "\finite(UNIV::'a set)" and - BIJ: "inj(g::'b \ 'a)" -shows "(\(f::('b + 'a) => 'a). bij f) \ (\(h::('a + 'b) => 'a). bij h)" -using assms Plus_infinite_bij_betw[of "UNIV::'a set" g "UNIV::'b set"] -by auto + assumes "\finite(UNIV::'a set)" and "inj(g::'b \ 'a)" + shows "(\(f::('b + 'a) => 'a). bij f) \ (\(h::('a + 'b) => 'a). bij h)" + using assms Plus_infinite_bij_betw[of "UNIV::'a set" g "UNIV::'b set"] + by auto lemma card_of_Un_infinite: -assumes INF: "\finite A" and LEQ: "|B| \o |A|" -shows "|A \ B| =o |A| \ |B \ A| =o |A|" -proof- - have "|A \ B| \o |A <+> B|" by (rule card_of_Un_Plus_ordLeq) - moreover have "|A <+> B| =o |A|" - using assms by (metis card_of_Plus_infinite) - ultimately have "|A \ B| \o |A|" using ordLeq_ordIso_trans by blast - hence "|A \ B| =o |A|" using card_of_Un1 ordIso_iff_ordLeq by blast - thus ?thesis using Un_commute[of B A] by auto -qed + assumes INF: "\finite A" and LEQ: "|B| \o |A|" + shows "|A \ B| =o |A| \ |B \ A| =o |A|" + by (simp add: INF LEQ card_of_Un_ordLeq_infinite_Field ordIso_iff_ordLeq) lemma card_of_Un_infinite_simps[simp]: -"\\finite A; |B| \o |A| \ \ |A \ B| =o |A|" -"\\finite A; |B| \o |A| \ \ |B \ A| =o |A|" -using card_of_Un_infinite by auto + "\\finite A; |B| \o |A| \ \ |A \ B| =o |A|" + "\\finite A; |B| \o |A| \ \ |B \ A| =o |A|" + using card_of_Un_infinite by auto lemma card_of_Un_diff_infinite: -assumes INF: "\finite A" and LESS: "|B| finite A" and LESS: "|B| B| =o |A|" - using assms ordLeq_iff_ordLess_or_ordIso card_of_Un_infinite by blast + using assms ordLeq_iff_ordLess_or_ordIso card_of_Un_infinite by blast moreover have "C \ B = A \ B" unfolding C_def by auto ultimately have 1: "|C \ B| =o |A|" by auto - (* *) + (* *) {assume *: "|C| \o |B|" - moreover - {assume **: "finite B" - hence "finite C" - using card_of_ordLeq_finite * by blast - hence False using ** INF card_of_ordIso_finite 1 by blast - } - hence "\finite B" by auto - ultimately have False - using card_of_Un_infinite 1 ordIso_equivalence(1,3) LESS not_ordLess_ordIso by metis + moreover + {assume **: "finite B" + hence "finite C" + using card_of_ordLeq_finite * by blast + hence False using ** INF card_of_ordIso_finite 1 by blast + } + hence "\finite B" by auto + ultimately have False + using card_of_Un_infinite 1 ordIso_equivalence(1,3) LESS not_ordLess_ordIso by metis } hence 2: "|B| \o |C|" using card_of_Well_order ordLeq_total by blast {assume *: "finite C" @@ -507,155 +445,121 @@ } hence "\finite C" by auto hence "|C| =o |A|" - using card_of_Un_infinite 1 2 ordIso_equivalence(1,3) by metis + using card_of_Un_infinite 1 2 ordIso_equivalence(1,3) by metis thus ?thesis unfolding C_def . qed corollary Card_order_Un_infinite: -assumes INF: "\finite(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" + assumes INF: "\finite(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 (auto simp add: card_of_Un_infinite) + using assms by (auto simp: card_of_Un_infinite) thus ?thesis - using assms card_of_Field_ordIso[of r] - ordIso_transitive[of "|Field r \ Field p|"] - ordIso_transitive[of _ "|Field r|"] by blast + using assms card_of_Field_ordIso[of r] + ordIso_transitive[of "|Field r \ Field p|"] + ordIso_transitive[of _ "|Field r|"] by blast qed corollary subset_ordLeq_diff_infinite: -assumes INF: "\finite B" and SUB: "A \ B" and LESS: "|A| finite (B - A)" -using assms card_of_Un_diff_infinite card_of_ordIso_finite by blast + assumes INF: "\finite B" and SUB: "A \ B" and LESS: "|A| finite (B - A)" + using assms card_of_Un_diff_infinite card_of_ordIso_finite by blast lemma card_of_Times_ordLess_infinite[simp]: -assumes INF: "\finite C" and - LESS1: "|A| B| finite C" and + LESS1: "|A| B| B = {}") assume Case1: "A = {} \ B = {}" hence "A \ B = {}" by blast moreover have "C \ {}" using - LESS1 card_of_empty5 by blast - ultimately show ?thesis by(auto simp add: card_of_empty4) + LESS1 card_of_empty5 by blast + ultimately show ?thesis by(auto simp: card_of_empty4) next assume Case2: "\(A = {} \ B = {})" {assume *: "|C| \o |A \ B|" - hence "\finite (A \ B)" using INF card_of_ordLeq_finite by blast - hence 1: "\finite A \ \finite B" using finite_cartesian_product by blast - {assume Case21: "|A| \o |B|" - hence "\finite B" using 1 card_of_ordLeq_finite by blast - hence "|A \ B| =o |B|" using Case2 Case21 - by (auto simp add: card_of_Times_infinite) - hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast - } - moreover - {assume Case22: "|B| \o |A|" - hence "\finite A" using 1 card_of_ordLeq_finite by blast - hence "|A \ B| =o |A|" using Case2 Case22 - by (auto simp add: card_of_Times_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 + hence "\finite (A \ B)" using INF card_of_ordLeq_finite by blast + hence 1: "\finite A \ \finite B" using finite_cartesian_product by blast + {assume Case21: "|A| \o |B|" + hence "\finite B" using 1 card_of_ordLeq_finite by blast + hence "|A \ B| =o |B|" using Case2 Case21 + by (auto simp: card_of_Times_infinite) + hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast + } + moreover + {assume Case22: "|B| \o |A|" + hence "\finite A" using 1 card_of_ordLeq_finite by blast + hence "|A \ B| =o |A|" using Case2 Case22 + by (auto simp: card_of_Times_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 + card_of_Well_order[of "A \ B"] card_of_Well_order[of "C"] by auto qed lemma card_of_Times_ordLess_infinite_Field[simp]: -assumes INF: "\finite (Field r)" and r: "Card_order r" and - LESS1: "|A| B| finite (Field r)" and r: "Card_order r" and + LESS1: "|A| B| |?C| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast + ordIso_symmetric by blast hence "|A| B| finite C" and - LESS1: "|A| B| finite (Field r)" and r: "Card_order r" and - LESS1: "|A| |?C| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast - hence "|A| o s" and "finite (Field s)" -shows "finite (Field r)" -by (metis assms card_of_mono2 card_of_ordLeq_infinite) + assumes "r \o s" and "finite (Field s)" + shows "finite (Field r)" + by (metis assms card_of_mono2 card_of_ordLeq_infinite) lemma ordIso_finite_Field: -assumes "r =o s" -shows "finite (Field r) \ finite (Field s)" -by (metis assms ordIso_iff_ordLeq ordLeq_finite_Field) + assumes "r =o s" + shows "finite (Field r) \ finite (Field s)" + by (metis assms ordIso_iff_ordLeq ordLeq_finite_Field) subsection \Cardinals versus set operations involving infinite sets\ lemma finite_iff_cardOf_nat: -"finite A = ( |A| finite B" -shows "|A| finite B" + shows "|A| finite A" -shows "|insert a A| =o |A|" + assumes "\finite A" + shows "|insert a A| =o |A|" proof- have iA: "insert a A = A \ {a}" by simp show ?thesis - using infinite_imp_bij_betw2[OF assms] unfolding iA - by (metis bij_betw_inv card_of_ordIso) + using infinite_imp_bij_betw2[OF assms] unfolding iA + by (metis bij_betw_inv card_of_ordIso) qed lemma card_of_Un_singl_ordLess_infinite1: -assumes "\finite B" and "|A| finite B" and "|A| finite B" -shows "( |A| finite B" + shows "|A| |{a} Un A| Cardinals versus lists\ @@ -664,48 +568,39 @@ proofs of facts concerning the cardinality of \List\ :\ definition nlists :: "'a set \ nat \ 'a list set" -where "nlists A n \ {l. set l \ A \ length l = n}" - -lemma lists_def2: "lists A = {l. set l \ A}" -using in_listsI by blast + where "nlists A n \ {l. set l \ A \ length l = n}" lemma lists_UNION_nlists: "lists A = (\n. nlists A n)" -unfolding lists_def2 nlists_def by blast + unfolding lists_eq_set nlists_def by blast lemma card_of_lists: "|A| \o |lists A|" proof- let ?h = "\ a. [a]" have "inj_on ?h A \ ?h ` A \ lists A" - unfolding inj_on_def lists_def2 by auto + unfolding inj_on_def lists_eq_set by auto thus ?thesis by (metis card_of_ordLeq) qed lemma nlists_0: "nlists A 0 = {[]}" -unfolding nlists_def by auto + unfolding nlists_def by auto lemma nlists_not_empty: -assumes "A \ {}" -shows "nlists A n \ {}" -proof(induct n, simp add: nlists_0) - fix n assume "nlists A n \ {}" + assumes "A \ {}" + shows "nlists A n \ {}" +proof (induction n) + case (Suc n) then obtain a and l where "a \ A \ l \ nlists A n" using assms by auto hence "a # l \ nlists A (Suc n)" unfolding nlists_def by auto - thus "nlists A (Suc n) \ {}" by auto -qed - -lemma Nil_in_lists: "[] \ lists A" -unfolding lists_def2 by auto - -lemma lists_not_empty: "lists A \ {}" -using Nil_in_lists by blast + then show ?case by auto +qed (simp add: nlists_0) lemma card_of_nlists_Succ: "|nlists A (Suc n)| =o |A \ (nlists A n)|" proof- let ?B = "A \ (nlists A n)" let ?h = "\(a,l). a # l" have "inj_on ?h ?B \ ?h ` ?B \ nlists A (Suc n)" - unfolding inj_on_def nlists_def by auto + unfolding inj_on_def nlists_def by auto moreover have "nlists A (Suc n) \ ?h ` ?B" - proof(auto) + proof clarify fix l assume "l \ nlists A (Suc n)" hence 1: "length l = Suc n \ set l \ A" unfolding nlists_def by auto then obtain a and l' where 2: "l = a # l'" by (auto simp: length_Suc_conv) @@ -713,131 +608,113 @@ thus "l \ ?h ` ?B" using 2 unfolding nlists_def by auto qed ultimately have "bij_betw ?h ?B (nlists A (Suc n))" - unfolding bij_betw_def by auto + unfolding bij_betw_def by auto thus ?thesis using card_of_ordIso ordIso_symmetric by blast qed lemma card_of_nlists_infinite: -assumes "\finite A" -shows "|nlists A n| \o |A|" -proof(induct n) + assumes "\finite A" + shows "|nlists A n| \o |A|" +proof(induction n) + case 0 have "A \ {}" using assms by auto - thus "|nlists A 0| \o |A|" by (simp add: nlists_0) + then show ?case + by (simp add: nlists_0) next - fix n assume IH: "|nlists A n| \o |A|" + case (Suc n) have "|nlists A (Suc n)| =o |A \ (nlists A n)|" - using card_of_nlists_Succ by blast + using card_of_nlists_Succ by blast moreover - {have "nlists A n \ {}" using assms nlists_not_empty[of A] by blast - hence "|A \ (nlists A n)| =o |A|" - using assms IH by (auto simp add: card_of_Times_infinite) - } - ultimately show "|nlists A (Suc n)| \o |A|" - using ordIso_transitive ordIso_iff_ordLeq by blast + have "nlists A n \ {}" using assms nlists_not_empty[of A] by blast + hence "|A \ (nlists A n)| =o |A|" + using Suc assms by auto + ultimately show ?case + using ordIso_transitive ordIso_iff_ordLeq by blast qed + lemma Card_order_lists: "Card_order r \ r \o |lists(Field r) |" -using card_of_lists card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast + using card_of_lists card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast lemma Union_set_lists: "\(set ` (lists A)) = A" - unfolding lists_def2 -proof(auto) - fix a assume "a \ A" +proof - + { fix a assume "a \ A" hence "set [a] \ A \ a \ set [a]" by auto - thus "\l. set l \ A \ a \ set l" by blast + hence "\l. set l \ A \ a \ set l" by blast } + then show ?thesis by force qed lemma inj_on_map_lists: -assumes "inj_on f A" -shows "inj_on (map f) (lists A)" -using assms Union_set_lists[of A] inj_on_mapI[of f "lists A"] by auto + assumes "inj_on f A" + shows "inj_on (map f) (lists A)" + using assms Union_set_lists[of A] inj_on_mapI[of f "lists A"] by auto lemma map_lists_mono: -assumes "f ` A \ B" -shows "(map f) ` (lists A) \ lists B" -using assms unfolding lists_def2 by (auto, blast) (* lethal combination of methods :) *) + assumes "f ` A \ B" + shows "(map f) ` (lists A) \ lists B" + using assms by force lemma map_lists_surjective: -assumes "f ` A = B" -shows "(map f) ` (lists A) = lists B" -using assms unfolding lists_def2 -proof (auto, blast) - fix l' assume *: "set l' \ f ` A" - have "set l' \ f ` A \ l' \ map f ` {l. set l \ A}" - proof(induct l', auto) - fix l a - assume "a \ A" and "set l \ A" and - IH: "f ` (set l) \ f ` A" - hence "set (a # l) \ A" by auto - hence "map f (a # l) \ map f ` {l. set l \ A}" by blast - thus "f a # map f l \ map f ` {l. set l \ A}" by auto - qed - thus "l' \ map f ` {l. set l \ A}" using * by auto -qed + assumes "f ` A = B" + shows "(map f) ` (lists A) = lists B" + by (metis assms lists_image) lemma bij_betw_map_lists: -assumes "bij_betw f A B" -shows "bij_betw (map f) (lists A) (lists B)" -using assms unfolding bij_betw_def -by(auto simp add: inj_on_map_lists map_lists_surjective) + assumes "bij_betw f A B" + shows "bij_betw (map f) (lists A) (lists B)" + using assms unfolding bij_betw_def + by(auto simp: inj_on_map_lists map_lists_surjective) lemma card_of_lists_mono[simp]: -assumes "|A| \o |B|" -shows "|lists A| \o |lists B|" + assumes "|A| \o |B|" + shows "|lists A| \o |lists B|" proof- obtain f where "inj_on f A \ f ` A \ B" - using assms card_of_ordLeq[of A B] by auto + using assms card_of_ordLeq[of A B] by auto hence "inj_on (map f) (lists A) \ (map f) ` (lists A) \ (lists B)" - by (auto simp add: inj_on_map_lists map_lists_mono) + by (auto simp: inj_on_map_lists map_lists_mono) thus ?thesis using card_of_ordLeq[of "lists A"] by metis qed lemma ordIso_lists_mono: -assumes "r \o r'" -shows "|lists(Field r)| \o |lists(Field r')|" -using assms card_of_mono2 card_of_lists_mono by blast + assumes "r \o r'" + shows "|lists(Field r)| \o |lists(Field r')|" + using assms card_of_mono2 card_of_lists_mono by blast lemma card_of_lists_cong[simp]: -assumes "|A| =o |B|" -shows "|lists A| =o |lists B|" -proof- - obtain f where "bij_betw f A B" - using assms card_of_ordIso[of A B] by auto - hence "bij_betw (map f) (lists A) (lists B)" - by (auto simp add: bij_betw_map_lists) - thus ?thesis using card_of_ordIso[of "lists A"] by auto -qed + assumes "|A| =o |B|" + shows "|lists A| =o |lists B|" + by (meson assms card_of_lists_mono ordIso_iff_ordLeq) lemma card_of_lists_infinite[simp]: -assumes "\finite A" -shows "|lists A| =o |A|" + assumes "\finite A" + shows "|lists A| =o |A|" proof- have "|lists A| \o |A|" unfolding lists_UNION_nlists - by (rule card_of_UNION_ordLeq_infinite[OF assms _ ballI[OF card_of_nlists_infinite[OF assms]]]) - (metis infinite_iff_card_of_nat assms) + by (rule card_of_UNION_ordLeq_infinite[OF assms _ ballI[OF card_of_nlists_infinite[OF assms]]]) + (metis infinite_iff_card_of_nat assms) thus ?thesis using card_of_lists ordIso_iff_ordLeq by blast qed lemma Card_order_lists_infinite: -assumes "Card_order r" and "\finite(Field r)" -shows "|lists(Field r)| =o r" -using assms card_of_lists_infinite card_of_Field_ordIso ordIso_transitive by blast + assumes "Card_order r" and "\finite(Field r)" + shows "|lists(Field r)| =o r" + using assms card_of_lists_infinite card_of_Field_ordIso ordIso_transitive by blast lemma ordIso_lists_cong: -assumes "r =o r'" -shows "|lists(Field r)| =o |lists(Field r')|" -using assms card_of_cong card_of_lists_cong by blast + assumes "r =o r'" + shows "|lists(Field r)| =o |lists(Field r')|" + using assms card_of_cong card_of_lists_cong by blast corollary lists_infinite_bij_betw: -assumes "\finite A" -shows "\f. bij_betw f (lists A) A" -using assms card_of_lists_infinite card_of_ordIso by blast + assumes "\finite A" + shows "\f. bij_betw f (lists A) A" + using assms card_of_lists_infinite card_of_ordIso by blast corollary lists_infinite_bij_betw_types: -assumes "\finite(UNIV :: 'a set)" -shows "\(f::'a list \ 'a). bij f" -using assms assms lists_infinite_bij_betw[of "UNIV::'a set"] -using lists_UNIV by auto + assumes "\finite(UNIV :: 'a set)" + shows "\(f::'a list \ 'a). bij f" + using assms lists_infinite_bij_betw by fastforce subsection \Cardinals versus the finite powerset operator\ @@ -846,82 +723,72 @@ proof- let ?h = "\ a. {a}" have "inj_on ?h A \ ?h ` A \ Fpow A" - unfolding inj_on_def Fpow_def by auto + unfolding inj_on_def Fpow_def by auto thus ?thesis using card_of_ordLeq by metis qed lemma Card_order_Fpow: "Card_order r \ r \o |Fpow(Field r) |" -using card_of_Fpow card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast + using card_of_Fpow card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast lemma image_Fpow_surjective: -assumes "f ` A = B" -shows "(image f) ` (Fpow A) = Fpow B" -using assms proof(unfold Fpow_def, auto) - fix Y assume *: "Y \ f ` A" and **: "finite Y" - hence "\b \ Y. \a. a \ A \ f a = b" by auto - with bchoice[of Y "\b a. a \ A \ f a = b"] - obtain g where 1: "\b \ Y. g b \ A \ f(g b) = b" by blast - obtain X where X_def: "X = g ` Y" by blast - have "f ` X = Y \ X \ A \ finite X" - by(unfold X_def, force simp add: ** 1) - thus "Y \ (image f) ` {X. X \ A \ finite X}" by auto + assumes "f ` A = B" + shows "(image f) ` (Fpow A) = Fpow B" +proof - + have "\C. \C \ f ` A; finite C\ \ C \ (`) f ` {X. X \ A \ finite X}" + by (simp add: finite_subset_image image_iff) + then show ?thesis + using assms by (force simp add: Fpow_def) qed lemma bij_betw_image_Fpow: -assumes "bij_betw f A B" -shows "bij_betw (image f) (Fpow A) (Fpow B)" -using assms unfolding bij_betw_def -by (auto simp add: inj_on_image_Fpow image_Fpow_surjective) + assumes "bij_betw f A B" + shows "bij_betw (image f) (Fpow A) (Fpow B)" + using assms unfolding bij_betw_def + by (auto simp: inj_on_image_Fpow image_Fpow_surjective) lemma card_of_Fpow_mono[simp]: -assumes "|A| \o |B|" -shows "|Fpow A| \o |Fpow B|" + assumes "|A| \o |B|" + shows "|Fpow A| \o |Fpow B|" proof- obtain f where "inj_on f A \ f ` A \ B" - using assms card_of_ordLeq[of A B] by auto + using assms card_of_ordLeq[of A B] by auto hence "inj_on (image f) (Fpow A) \ (image f) ` (Fpow A) \ (Fpow B)" - by (auto simp add: inj_on_image_Fpow image_Fpow_mono) + by (auto simp: inj_on_image_Fpow image_Fpow_mono) thus ?thesis using card_of_ordLeq[of "Fpow A"] by auto qed lemma ordIso_Fpow_mono: -assumes "r \o r'" -shows "|Fpow(Field r)| \o |Fpow(Field r')|" -using assms card_of_mono2 card_of_Fpow_mono by blast + assumes "r \o r'" + shows "|Fpow(Field r)| \o |Fpow(Field r')|" + using assms card_of_mono2 card_of_Fpow_mono by blast lemma card_of_Fpow_cong[simp]: -assumes "|A| =o |B|" -shows "|Fpow A| =o |Fpow B|" -proof- - obtain f where "bij_betw f A B" - using assms card_of_ordIso[of A B] by auto - hence "bij_betw (image f) (Fpow A) (Fpow B)" - by (auto simp add: bij_betw_image_Fpow) - thus ?thesis using card_of_ordIso[of "Fpow A"] by auto -qed + assumes "|A| =o |B|" + shows "|Fpow A| =o |Fpow B|" + by (meson assms card_of_Fpow_mono ordIso_iff_ordLeq) lemma ordIso_Fpow_cong: -assumes "r =o r'" -shows "|Fpow(Field r)| =o |Fpow(Field r')|" -using assms card_of_cong card_of_Fpow_cong by blast + assumes "r =o r'" + shows "|Fpow(Field r)| =o |Fpow(Field r')|" + using assms card_of_cong card_of_Fpow_cong by blast lemma card_of_Fpow_lists: "|Fpow A| \o |lists A|" proof- have "set ` (lists A) = Fpow A" - unfolding lists_def2 Fpow_def using finite_list finite_set by blast + unfolding lists_eq_set Fpow_def using finite_list finite_set by blast thus ?thesis using card_of_ordLeq2[of "Fpow A"] Fpow_not_empty[of A] by blast qed lemma card_of_Fpow_infinite[simp]: -assumes "\finite A" -shows "|Fpow A| =o |A|" -using assms card_of_Fpow_lists card_of_lists_infinite card_of_Fpow - ordLeq_ordIso_trans ordIso_iff_ordLeq by blast + assumes "\finite A" + shows "|Fpow A| =o |A|" + using assms card_of_Fpow_lists card_of_lists_infinite card_of_Fpow + ordLeq_ordIso_trans ordIso_iff_ordLeq by blast corollary Fpow_infinite_bij_betw: -assumes "\finite A" -shows "\f. bij_betw f (Fpow A) A" -using assms card_of_Fpow_infinite card_of_ordIso by blast + assumes "\finite A" + shows "\f. bij_betw f (Fpow A) A" + using assms card_of_Fpow_infinite card_of_ordIso by blast subsection \The cardinal $\omega$ and the finite cardinals\ @@ -929,108 +796,120 @@ subsubsection \First as well-orders\ lemma Field_natLess: "Field natLess = (UNIV::nat set)" -by(unfold Field_def natLess_def, auto) + by (auto simp: Field_def natLess_def) lemma natLeq_well_order_on: "well_order_on UNIV natLeq" -using natLeq_Well_order Field_natLeq by auto + using natLeq_Well_order Field_natLeq by auto lemma natLeq_wo_rel: "wo_rel natLeq" -unfolding wo_rel_def using natLeq_Well_order . + unfolding wo_rel_def using natLeq_Well_order . lemma natLeq_ofilter_less: "ofilter natLeq {0 ..< n}" -by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def, - simp add: Field_natLeq, unfold under_def natLeq_def, auto) +proof - + have "\t Field natLeq" + by (simp add: Field_natLeq) + moreover have "\xt. (t, x) \ natLeq \ t < n" + by (simp add: natLeq_def) + ultimately show ?thesis + by (auto simp: natLeq_wo_rel wo_rel.ofilter_def under_def) +qed lemma natLeq_ofilter_leq: "ofilter natLeq {0 .. n}" -by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def, - simp add: Field_natLeq, unfold under_def natLeq_def, auto) + by (metis (no_types) atLeastLessThanSuc_atLeastAtMost natLeq_ofilter_less) lemma natLeq_UNIV_ofilter: "wo_rel.ofilter natLeq UNIV" -using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto + using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto lemma closed_nat_set_iff: -assumes "\(m::nat) n. n \ A \ m \ n \ m \ A" -shows "A = UNIV \ (\n. A = {0 ..< n})" + 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 + 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 + then have "A = {0 ..< n}" + proof(auto simp: 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 natLeq_ofilter_iff: -"ofilter natLeq A = (A = UNIV \ (\n. A = {0 ..< n}))" + "ofilter natLeq A = (A = UNIV \ (\n. A = {0 ..< n}))" proof(rule iffI) assume "ofilter natLeq A" hence "\m n. n \ A \ m \ n \ m \ A" using natLeq_wo_rel - by(auto simp add: natLeq_def wo_rel.ofilter_def under_def) + by(auto simp: natLeq_def wo_rel.ofilter_def under_def) thus "A = UNIV \ (\n. A = {0 ..< n})" using closed_nat_set_iff by blast next assume "A = UNIV \ (\n. A = {0 ..< n})" thus "ofilter natLeq A" - by(auto simp add: natLeq_ofilter_less natLeq_UNIV_ofilter) + by(auto simp: natLeq_ofilter_less natLeq_UNIV_ofilter) qed lemma natLeq_under_leq: "under natLeq n = {0 .. n}" -unfolding under_def natLeq_def by auto + unfolding under_def natLeq_def by auto lemma natLeq_on_ofilter_less_eq: -"n \ m \ wo_rel.ofilter (natLeq_on m) {0 ..< n}" -apply (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def) -apply (simp add: Field_natLeq_on) -by (auto simp add: under_def) + "n \ m \ wo_rel.ofilter (natLeq_on m) {0 ..< n}" + by (auto simp: natLeq_on_wo_rel wo_rel.ofilter_def Field_natLeq_on under_def) lemma natLeq_on_ofilter_iff: -"wo_rel.ofilter (natLeq_on m) A = (\n \ m. A = {0 ..< n})" + "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 under_def) + using * by(fastforce simp add: natLeq_on_wo_rel wo_rel.ofilter_def 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) + thus "wo_rel.ofilter (natLeq_on m) A" by (auto simp: natLeq_on_ofilter_less_eq) qed corollary natLeq_on_ofilter: -"ofilter(natLeq_on n) {0 ..< n}" -by (auto simp add: natLeq_on_ofilter_less_eq) + "ofilter(natLeq_on n) {0 ..< n}" + by (auto simp: natLeq_on_ofilter_less_eq) lemma natLeq_on_ofilter_less: -"n < m \ ofilter (natLeq_on m) {0 .. n}" -by(auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def, - simp add: Field_natLeq_on, unfold under_def, auto) + assumes "n < m" shows "ofilter (natLeq_on m) {0 .. n}" +proof - + have "Suc n \ m" + using assms by simp + then show ?thesis + using natLeq_on_ofilter_iff by auto +qed lemma natLeq_on_ordLess_natLeq: "natLeq_on n n. well_order_on {na. na < n} (natLeq_on n)" + using Field_natLeq_on natLeq_on_Well_order by presburger + ultimately show ?thesis + by (simp add: Field_natLeq Field_natLeq_on finite_ordLess_infinite) +qed lemma natLeq_on_injective: -"natLeq_on m = natLeq_on n \ m = n" -using Field_natLeq_on[of m] Field_natLeq_on[of n] - atLeastLessThan_injective[of m n, unfolded atLeastLessThan_def] by blast + "natLeq_on m = natLeq_on n \ m = n" + using Field_natLeq_on[of m] Field_natLeq_on[of n] + atLeastLessThan_injective[of m n, unfolded atLeastLessThan_def] by blast lemma natLeq_on_injective_ordIso: -"(natLeq_on m =o natLeq_on n) = (m = n)" -proof(auto simp add: natLeq_on_Well_order ordIso_reflexive) + "(natLeq_on m =o natLeq_on n) = (m = n)" +proof(auto simp: natLeq_on_Well_order ordIso_reflexive) assume "natLeq_on m =o natLeq_on n" then obtain f where "bij_betw f {x. xThen as cardinals\ lemma ordIso_natLeq_infinite1: -"|A| =o natLeq \ \finite A" -using ordIso_symmetric ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast + "|A| =o natLeq \ \finite A" + by (meson finite_iff_ordLess_natLeq not_ordLess_ordIso) lemma ordIso_natLeq_infinite2: -"natLeq =o |A| \ \finite A" -using ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast + "natLeq =o |A| \ \finite A" + using ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast lemma ordIso_natLeq_on_imp_finite: -"|A| =o natLeq_on n \ finite A" -unfolding ordIso_def iso_def[abs_def] -by (auto simp: Field_natLeq_on bij_betw_finite) + "|A| =o natLeq_on n \ finite A" + unfolding ordIso_def iso_def[abs_def] + by (auto simp: Field_natLeq_on bij_betw_finite) 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 {x. x < n} r" - thus "natLeq_on n \o r" - using finite_atLeastLessThan natLeq_on_well_order_on +proof - + { fix r assume "well_order_on {x. x < n} r" + hence "natLeq_on n \o r" + using finite_atLeastLessThan natLeq_on_well_order_on finite_well_order_on_ordIso ordIso_iff_ordLeq by blast + } + then show ?thesis + unfolding card_order_on_def using Field_natLeq_on natLeq_on_Well_order by presburger 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 + "|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 -unfolding atLeast_0 atLeastLessThan_def lessThan_def Int_UNIV_left by auto + "|{0 ..< n}| =o natLeq_on n" + using Field_natLeq_on card_of_Field_natLeq_on + unfolding atLeast_0 atLeastLessThan_def lessThan_def Int_UNIV_left by auto lemma natLeq_on_ordLeq_less_eq: -"((natLeq_on m) \o (natLeq_on n)) = (m \ n)" + "((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 {x. x < m} \ f ` {x. x < m} \ {x. x < n}" - unfolding ordLeq_def using - embed_inj_on[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on] - embed_Field Field_natLeq_on by blast + unfolding ordLeq_def using + embed_inj_on[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on] + embed_Field Field_natLeq_on by blast thus "m \ n" using atLeastLessThan_less_eq2 unfolding atLeast_0 atLeastLessThan_def lessThan_def Int_UNIV_left by blast next @@ -1086,141 +967,126 @@ 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 + using card_of_less ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast qed lemma natLeq_on_ordLeq_less: -"((natLeq_on m) o natLeq_on n" -shows "finite A" + assumes "|A| \o natLeq_on n" + shows "finite A" proof- have "|A| \o |{0 ..< n}|" - using assms card_of_less ordIso_symmetric ordLeq_ordIso_trans by blast - thus ?thesis by (auto simp add: card_of_ordLeq_finite) + using assms card_of_less ordIso_symmetric ordLeq_ordIso_trans by blast + thus ?thesis by (auto simp: card_of_ordLeq_finite) qed subsubsection \"Backward compatibility" with the numeric cardinal operator for finite sets\ lemma finite_card_of_iff_card2: -assumes FIN: "finite A" and FIN': "finite B" -shows "( |A| \o |B| ) = (card A \ card B)" -using assms card_of_ordLeq[of A B] inj_on_iff_card_le[of A B] by blast + assumes FIN: "finite A" and FIN': "finite B" + shows "( |A| \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)" + 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 + 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) + "finite A = (\n. |A| =o natLeq_on n)" + using finite_imp_card_of_natLeq_on[of A] + by(auto simp: ordIso_natLeq_on_imp_finite) lemma finite_card_of_iff_card: -assumes FIN: "finite A" and FIN': "finite B" -shows "( |A| =o |B| ) = (card A = card B)" -using assms card_of_ordIso[of A B] bij_betw_iff_card[of A B] by blast + assumes FIN: "finite A" and FIN': "finite B" + shows "( |A| =o |B| ) = (card A = card B)" + using assms card_of_ordIso[of A B] bij_betw_iff_card[of A B] by blast lemma finite_card_of_iff_card3: -assumes FIN: "finite A" and FIN': "finite B" -shows "( |A| o |A| ))" by simp - also have "... = (~ (card B \ card A))" - using assms by(simp add: finite_card_of_iff_card2) - also have "... = (card A < card B)" by auto + also have "\ = (~ (card B \ card A))" + using assms by(simp add: finite_card_of_iff_card2) + also have "\ = (card A < card B)" by auto finally show ?thesis . qed lemma card_Field_natLeq_on: -"card(Field(natLeq_on n)) = n" -using Field_natLeq_on card_atLeastLessThan by auto + "card(Field(natLeq_on n)) = n" + using Field_natLeq_on card_atLeastLessThan by auto subsection \The successor of a cardinal\ lemma embed_implies_ordIso_Restr: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and EMB: "embed r' r f" -shows "r' =o Restr r (f ` (Field r'))" -using assms embed_implies_iso_Restr Well_order_Restr unfolding ordIso_def by blast - -lemma cardSuc_Well_order[simp]: -"Card_order r \ Well_order(cardSuc r)" -using cardSuc_Card_order unfolding card_order_on_def by blast - -lemma Field_cardSuc_not_empty: -assumes "Card_order r" -shows "Field (cardSuc r) \ {}" -proof - assume "Field(cardSuc r) = {}" - hence "|Field(cardSuc r)| \o r" using assms Card_order_empty[of r] by auto - hence "cardSuc r \o r" using assms card_of_Field_ordIso - cardSuc_Card_order ordIso_symmetric ordIso_ordLeq_trans by blast - thus False using cardSuc_greater not_ordLess_ordLeq assms by blast -qed + assumes WELL: "Well_order r" and WELL': "Well_order r'" and EMB: "embed r' r f" + shows "r' =o Restr r (f ` (Field r'))" + using assms embed_implies_iso_Restr Well_order_Restr unfolding ordIso_def by blast lemma cardSuc_mono_ordLess[simp]: -assumes CARD: "Card_order r" and CARD': "Card_order r'" -shows "(cardSuc r Well_order r' \ Well_order(cardSuc r) \ Well_order(cardSuc r')" - using assms by auto + using assms by auto thus ?thesis - using not_ordLeq_iff_ordLess not_ordLeq_iff_ordLess[of r r'] - using cardSuc_mono_ordLeq[of r' r] assms by blast + using not_ordLeq_iff_ordLess not_ordLeq_iff_ordLess[of r r'] + using cardSuc_mono_ordLeq[of r' r] assms by blast qed lemma cardSuc_natLeq_on_Suc: -"cardSuc(natLeq_on n) =o natLeq_on(Suc n)" + "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: *) + 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 + 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 + 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 atLeast_0 atLeastLessThan_def lessThan_def by simp + unfolding r_def p_def Field_natLeq_on atLeast_0 atLeastLessThan_def lessThan_def 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 + 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 WELL embed_inj_on[OF _ 1] unfolding bij_betw_def by blast - 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 + then obtain f where 0: "embedS r' p f" unfolding ordLess_def by force + let ?q = "Restr p (f ` Field r')" + have 1: "embed r' p f" using 0 unfolding embedS_def by force + hence 2: "f ` Field r' < {0..<(Suc n)}" + using WELL FIELD 0 by (auto simp: embedS_iff) + have "wo_rel.ofilter p (f ` Field r')" using embed_Field_ofilter 1 WELL by blast + then obtain m where "m \ Suc n" and 3: "f ` (Field r') = {0.. n" using 2 by force + (* *) + have "bij_betw f (Field r') (f ` (Field r'))" + using WELL embed_inj_on[OF _ 1] unfolding bij_betw_def by blast + 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 qed @@ -1228,129 +1094,104 @@ qed lemma card_of_Plus_ordLeq_infinite[simp]: -assumes C: "\finite C" and A: "|A| \o |C|" and B: "|B| \o |C|" -shows "|A <+> B| \o |C|" -proof- - let ?r = "cardSuc |C|" - have "Card_order ?r \ \finite (Field ?r)" using assms by simp - moreover have "|A| B| finite C" and "|A| \o |C|" and "|B| \o |C|" + shows "|A <+> B| \o |C|" + by (simp add: assms) lemma card_of_Un_ordLeq_infinite[simp]: -assumes C: "\finite C" and A: "|A| \o |C|" and B: "|B| \o |C|" -shows "|A Un B| \o |C|" -using assms card_of_Plus_ordLeq_infinite card_of_Un_Plus_ordLeq -ordLeq_transitive by metis + assumes "\finite C" and "|A| \o |C|" and "|B| \o |C|" + shows "|A Un B| \o |C|" + using assms card_of_Plus_ordLeq_infinite card_of_Un_Plus_ordLeq ordLeq_transitive by metis subsection \Others\ lemma under_mono[simp]: -assumes "Well_order r" and "(i,j) \ r" -shows "under r i \ under r j" -using assms unfolding under_def order_on_defs -trans_def by blast + assumes "Well_order r" and "(i,j) \ r" + shows "under r i \ under r j" + using assms unfolding under_def order_on_defs trans_def by blast lemma underS_under: -assumes "i \ Field r" -shows "underS r i = under r i - {i}" -using assms unfolding underS_def under_def by auto + assumes "i \ Field r" + shows "underS r i = under r i - {i}" + using assms unfolding underS_def under_def by auto lemma relChain_under: -assumes "Well_order r" -shows "relChain r (\ i. under r i)" -using assms unfolding relChain_def by auto + assumes "Well_order r" + shows "relChain r (\ i. under r i)" + using assms unfolding relChain_def by auto lemma card_of_infinite_diff_finite: -assumes "\finite A" and "finite B" -shows "|A - B| =o |A|" -by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2) + assumes "\finite A" and "finite B" + shows "|A - B| =o |A|" + by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2) lemma infinite_card_of_diff_singl: -assumes "\finite A" -shows "|A - {a}| =o |A|" -by (metis assms card_of_infinite_diff_finite finite.emptyI finite_insert) + assumes "\finite A" + shows "|A - {a}| =o |A|" + by (metis assms card_of_infinite_diff_finite finite.emptyI finite_insert) lemma card_of_vimage: -assumes "B \ range f" -shows "|B| \o |f -` B|" -apply(rule surj_imp_ordLeq[of _ f]) -using assms by (metis Int_absorb2 image_vimage_eq order_refl) + assumes "B \ range f" + shows "|B| \o |f -` B|" + by (metis Int_absorb2 assms image_vimage_eq order_refl surj_imp_ordLeq) lemma surj_card_of_vimage: -assumes "surj f" -shows "|B| \o |f -` B|" -by (metis assms card_of_vimage subset_UNIV) - -lemma infinite_Pow: -assumes "\ finite A" -shows "\ finite (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 + assumes "surj f" + shows "|B| \o |f -` B|" + by (metis assms card_of_vimage subset_UNIV) (* bounded powerset *) definition Bpow where -"Bpow r A \ {X . X \ A \ |X| \o r}" + "Bpow r A \ {X . X \ A \ |X| \o r}" lemma Bpow_empty[simp]: -assumes "Card_order r" -shows "Bpow r {} = {{}}" -using assms unfolding Bpow_def by auto + assumes "Card_order r" + shows "Bpow r {} = {{}}" + using assms unfolding Bpow_def by auto lemma singl_in_Bpow: -assumes rc: "Card_order r" -and r: "Field r \ {}" and a: "a \ A" -shows "{a} \ Bpow r A" + assumes rc: "Card_order r" + and r: "Field r \ {}" and a: "a \ A" + shows "{a} \ Bpow r A" proof- have "|{a}| \o r" using r rc by auto thus ?thesis unfolding Bpow_def using a by auto qed lemma ordLeq_card_Bpow: -assumes rc: "Card_order r" and r: "Field r \ {}" -shows "|A| \o |Bpow r A|" + assumes rc: "Card_order r" and r: "Field r \ {}" + shows "|A| \o |Bpow r A|" proof- have "inj_on (\ a. {a}) A" unfolding inj_on_def by auto moreover have "(\ a. {a}) ` A \ Bpow r A" - using singl_in_Bpow[OF assms] by auto + using singl_in_Bpow[OF assms] by auto ultimately show ?thesis unfolding card_of_ordLeq[symmetric] by blast qed lemma infinite_Bpow: -assumes rc: "Card_order r" and r: "Field r \ {}" -and A: "\finite A" -shows "\finite (Bpow r A)" -using ordLeq_card_Bpow[OF rc r] -by (metis A card_of_ordLeq_infinite) + assumes rc: "Card_order r" and r: "Field r \ {}" + and A: "\finite A" + shows "\finite (Bpow r A)" + using ordLeq_card_Bpow[OF rc r] + by (metis A card_of_ordLeq_infinite) definition Func_option where - "Func_option A B \ + "Func_option A B \ {f. (\ a. f a \ None \ a \ A) \ (\ a \ A. case f a of Some b \ b \ B |None \ True)}" lemma card_of_Func_option_Func: -"|Func_option A B| =o |Func A B|" + "|Func_option A B| =o |Func A B|" proof (rule ordIso_symmetric, unfold card_of_ordIso[symmetric], intro exI) let ?F = "\ f a. if a \ A then Some (f a) else None" show "bij_betw ?F (Func A B) (Func_option A B)" - unfolding bij_betw_def unfolding inj_on_def proof(intro conjI ballI impI) + unfolding bij_betw_def unfolding inj_on_def proof(intro conjI ballI impI) fix f g assume f: "f \ Func A B" and g: "g \ Func A B" and eq: "?F f = ?F g" show "f = g" proof(rule ext) fix a show "f a = g a" - proof(cases "a \ A") - case True - have "Some (f a) = ?F f a" using True by auto - also have "... = ?F g a" using eq unfolding fun_eq_iff by(rule allE) - also have "... = Some (g a)" using True by auto - finally have "Some (f a) = Some (g a)" . - thus ?thesis by simp - qed(insert f g, unfold Func_def, auto) + by (smt (verit) Func_def eq f g mem_Collect_eq option.inject) qed next show "?F ` Func A B = Func_option A B" @@ -1358,11 +1199,11 @@ fix f assume f: "f \ Func_option A B" define g where [abs_def]: "g a = (case f a of Some b \ b | None \ undefined)" for a have "g \ Func A B" - using f unfolding g_def Func_def Func_option_def by force+ + using f unfolding g_def Func_def Func_option_def by force+ moreover have "f = ?F g" proof(rule ext) fix a show "f a = ?F g a" - using f unfolding Func_option_def g_def by (cases "a \ A") force+ + using f unfolding Func_option_def g_def by (cases "a \ A") force+ qed ultimately show "f \ ?F ` (Func A B)" by blast qed(unfold Func_def Func_option_def, auto) @@ -1371,16 +1212,16 @@ (* partial-function space: *) definition Pfunc where -"Pfunc A B \ + "Pfunc A B \ {f. (\a. f a \ None \ a \ A) \ (\a. case f a of None \ True | Some b \ b \ B)}" lemma Func_Pfunc: -"Func_option A B \ Pfunc A B" -unfolding Func_option_def Pfunc_def by auto + "Func_option A B \ Pfunc A B" + unfolding Func_option_def Pfunc_def by auto lemma Pfunc_Func_option: -"Pfunc A B = (\A' \ Pow A. Func_option A' B)" + "Pfunc A B = (\A' \ Pow A. Func_option A' B)" proof safe fix f assume f: "f \ Pfunc A B" show "f \ (\A'\Pow A. Func_option A' B)" @@ -1395,180 +1236,161 @@ qed lemma card_of_Func_mono: -fixes A1 A2 :: "'a set" and B :: "'b set" -assumes A12: "A1 \ A2" and B: "B \ {}" -shows "|Func A1 B| \o |Func A2 B|" + fixes A1 A2 :: "'a set" and B :: "'b set" + assumes A12: "A1 \ A2" and B: "B \ {}" + shows "|Func A1 B| \o |Func A2 B|" proof- obtain bb where bb: "bb \ B" using B by auto define F where [abs_def]: "F f1 a = (if a \ A2 then (if a \ A1 then f1 a else bb) else undefined)" for f1 :: "'a \ 'b" and a - show ?thesis unfolding card_of_ordLeq[symmetric] proof(intro exI[of _ F] conjI) - show "inj_on F (Func A1 B)" unfolding inj_on_def proof safe + show ?thesis unfolding card_of_ordLeq[symmetric] + proof(intro exI[of _ F] conjI) + show "inj_on F (Func A1 B)" unfolding inj_on_def + proof safe fix f g assume f: "f \ Func A1 B" and g: "g \ Func A1 B" and eq: "F f = F g" show "f = g" proof(rule ext) fix a show "f a = g a" - proof(cases "a \ A1") - case True - thus ?thesis using eq A12 unfolding F_def fun_eq_iff - by (elim allE[of _ a]) auto - qed(insert f g, unfold Func_def, fastforce) + by (smt (verit) A12 F_def Func_def eq f g mem_Collect_eq subsetD) qed qed qed(insert bb, unfold Func_def F_def, force) qed lemma card_of_Func_option_mono: -fixes A1 A2 :: "'a set" and B :: "'b set" -assumes A12: "A1 \ A2" and B: "B \ {}" -shows "|Func_option A1 B| \o |Func_option A2 B|" -by (metis card_of_Func_mono[OF A12 B] card_of_Func_option_Func - ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric) + fixes A1 A2 :: "'a set" and B :: "'b set" + assumes A12: "A1 \ A2" and B: "B \ {}" + shows "|Func_option A1 B| \o |Func_option A2 B|" + by (metis card_of_Func_mono[OF A12 B] card_of_Func_option_Func + ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric) lemma card_of_Pfunc_Pow_Func_option: -assumes "B \ {}" -shows "|Pfunc A B| \o |Pow A \ Func_option A B|" + assumes "B \ {}" + shows "|Pfunc A B| \o |Pow A \ Func_option A B|" proof- have "|Pfunc A B| =o |\A' \ Pow A. Func_option A' B|" (is "_ =o ?K") unfolding Pfunc_Func_option by(rule card_of_refl) also have "?K \o |Sigma (Pow A) (\ A'. Func_option A' B)|" using card_of_UNION_Sigma . also have "|Sigma (Pow A) (\ A'. Func_option A' B)| \o |Pow A \ Func_option A B|" - apply(rule card_of_Sigma_mono1) using card_of_Func_option_mono[OF _ assms] by auto + by (simp add: assms card_of_Func_option_mono card_of_Sigma_mono1) finally show ?thesis . qed lemma Bpow_ordLeq_Func_Field: -assumes rc: "Card_order r" and r: "Field r \ {}" and A: "\finite A" -shows "|Bpow r A| \o |Func (Field r) A|" + assumes rc: "Card_order r" and r: "Field r \ {}" and A: "\finite A" + shows "|Bpow r A| \o |Func (Field r) A|" proof- let ?F = "\ f. {x | x a. f a = x \ a \ Field r}" {fix X assume "X \ Bpow r A - {{}}" - hence XA: "X \ A" and "|X| \o r" - and X: "X \ {}" unfolding Bpow_def by auto - hence "|X| \o |Field r|" by (metis Field_card_of card_of_mono2) - then obtain F where 1: "X = F ` (Field r)" - using card_of_ordLeq2[OF X] by metis - define f where [abs_def]: "f i = (if i \ Field r then F i else undefined)" for i - have "\ f \ Func (Field r) A. X = ?F f" - apply (intro bexI[of _ f]) using 1 XA unfolding Func_def f_def by auto + hence XA: "X \ A" and "|X| \o r" + and X: "X \ {}" unfolding Bpow_def by auto + hence "|X| \o |Field r|" by (metis Field_card_of card_of_mono2) + then obtain F where 1: "X = F ` (Field r)" + using card_of_ordLeq2[OF X] by metis + define f where [abs_def]: "f i = (if i \ Field r then F i else undefined)" for i + have "\ f \ Func (Field r) A. X = ?F f" + apply (intro bexI[of _ f]) using 1 XA unfolding Func_def f_def by auto } hence "Bpow r A - {{}} \ ?F ` (Func (Field r) A)" by auto hence "|Bpow r A - {{}}| \o |Func (Field r) A|" - by (rule surj_imp_ordLeq) + by (rule surj_imp_ordLeq) moreover {have 2: "\finite (Bpow r A)" using infinite_Bpow[OF rc r A] . - have "|Bpow r A| =o |Bpow r A - {{}}|" - by (metis 2 infinite_card_of_diff_singl ordIso_symmetric) + have "|Bpow r A| =o |Bpow r A - {{}}|" + by (metis 2 infinite_card_of_diff_singl ordIso_symmetric) } ultimately show ?thesis by (metis ordIso_ordLeq_trans) qed lemma empty_in_Func[simp]: -"B \ {} \ (\x. undefined) \ Func {} B" -unfolding Func_def by auto + "B \ {} \ (\x. undefined) \ Func {} B" + by simp lemma Func_mono[simp]: -assumes "B1 \ B2" -shows "Func A B1 \ Func A B2" -using assms unfolding Func_def by force + assumes "B1 \ B2" + shows "Func A B1 \ Func A B2" + using assms unfolding Func_def by force lemma Pfunc_mono[simp]: -assumes "A1 \ A2" and "B1 \ B2" -shows "Pfunc A B1 \ Pfunc A B2" -using assms unfolding Pfunc_def -apply safe -apply (case_tac "x a", auto) -apply (metis in_mono option.simps(5)) -done + assumes "A1 \ A2" and "B1 \ B2" + shows "Pfunc A B1 \ Pfunc A B2" + using assms unfolding Pfunc_def + by (force split: option.split_asm option.split) lemma card_of_Func_UNIV_UNIV: -"|Func (UNIV::'a set) (UNIV::'b set)| =o |UNIV::('a \ 'b) set|" -using card_of_Func_UNIV[of "UNIV::'b set"] by auto + "|Func (UNIV::'a set) (UNIV::'b set)| =o |UNIV::('a \ 'b) set|" + using card_of_Func_UNIV[of "UNIV::'b set"] by auto lemma ordLeq_Func: -assumes "{b1,b2} \ B" "b1 \ b2" -shows "|A| \o |Func A B|" -unfolding card_of_ordLeq[symmetric] proof(intro exI conjI) - let ?F = "\ aa a. if a \ A then (if a = aa then b1 else b2) else undefined" + assumes "{b1,b2} \ B" "b1 \ b2" + shows "|A| \o |Func A B|" + unfolding card_of_ordLeq[symmetric] proof(intro exI conjI) + let ?F = "\x a. if a \ A then (if a = x then b1 else b2) else undefined" show "inj_on ?F A" using assms unfolding inj_on_def fun_eq_iff by auto show "?F ` A \ Func A B" using assms unfolding Func_def by auto qed lemma infinite_Func: -assumes A: "\finite A" and B: "{b1,b2} \ B" "b1 \ b2" -shows "\finite (Func A B)" -using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite) + assumes A: "\finite A" and B: "{b1,b2} \ B" "b1 \ b2" + shows "\finite (Func A B)" + using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite) subsection \Infinite cardinals are limit ordinals\ lemma card_order_infinite_isLimOrd: -assumes c: "Card_order r" and i: "\finite (Field r)" -shows "isLimOrd r" + assumes c: "Card_order r" and i: "\finite (Field r)" + shows "isLimOrd r" proof- have 0: "wo_rel r" and 00: "Well_order r" - using c unfolding card_order_on_def wo_rel_def by auto + using c unfolding card_order_on_def wo_rel_def by auto hence rr: "Refl r" by (metis wo_rel.REFL) - show ?thesis unfolding wo_rel.isLimOrd_def[OF 0] wo_rel.isSuccOrd_def[OF 0] proof safe - fix j assume j: "j \ Field r" and jm: "\i\Field r. (i, j) \ r" - define p where "p = Restr r (Field r - {j})" - have fp: "Field p = Field r - {j}" - unfolding p_def apply(rule Refl_Field_Restr2[OF rr]) by auto - have of: "ofilter r (Field p)" unfolding wo_rel.ofilter_def[OF 0] proof safe - fix a x assume a: "a \ Field p" and "x \ under r a" - hence x: "(x,a) \ r" "x \ Field r" unfolding under_def Field_def by auto - moreover have a: "a \ j" "a \ Field r" "(a,j) \ r" using a jm unfolding fp by auto - ultimately have "x \ j" using j jm by (metis 0 wo_rel.max2_def wo_rel.max2_equals1) - thus "x \ Field p" using x unfolding fp by auto - qed(unfold p_def Field_def, auto) - have "p Field r" and "\i\Field r. (i, j) \ r" + then show False + by (metis Card_order_trans c i infinite_Card_order_limit) qed qed lemma insert_Chain: -assumes "Refl r" "C \ Chains r" and "i \ Field r" and "\j. j \ C \ (j,i) \ r \ (i,j) \ r" -shows "insert i C \ Chains r" -using assms unfolding Chains_def by (auto dest: refl_onD) + assumes "Refl r" "C \ Chains r" and "i \ Field r" and "\j. j \ C \ (j,i) \ r \ (i,j) \ r" + shows "insert i C \ Chains r" + using assms unfolding Chains_def by (auto dest: refl_onD) lemma Collect_insert: "{R j |j. j \ insert j1 J} = insert (R j1) {R j |j. j \ J}" -by auto + by auto lemma Field_init_seg_of[simp]: -"Field init_seg_of = UNIV" -unfolding Field_def init_seg_of_def by auto + "Field init_seg_of = UNIV" + unfolding Field_def init_seg_of_def by auto lemma refl_init_seg_of[intro, simp]: "refl init_seg_of" -unfolding refl_on_def Field_def by auto + unfolding refl_on_def Field_def by auto lemma regularCard_all_ex: -assumes r: "Card_order r" "regularCard r" -and As: "\ i j b. b \ B \ (i,j) \ r \ P i b \ P j b" -and Bsub: "\ b \ B. \ i \ Field r. P i b" -and cardB: "|B| i \ Field r. \ b \ B. P i b" + assumes r: "Card_order r" "regularCard r" + and As: "\ i j b. b \ B \ (i,j) \ r \ P i b \ P j b" + and Bsub: "\ b \ B. \ i \ Field r. P i b" + and cardB: "|B| i \ Field r. \ b \ B. P i b" proof- let ?As = "\i. {b \ B. P i b}" have "\i \ Field r. B \ ?As i" - apply(rule regularCard_UNION) using assms unfolding relChain_def by auto + apply(rule regularCard_UNION) using assms unfolding relChain_def by auto thus ?thesis by auto qed lemma relChain_stabilize: -assumes rc: "relChain r As" and AsB: "(\i \ Field r. As i) \ B" and Br: "|B| finite (Field r)" and cr: "Card_order r" -shows "\ i \ Field r. As (succ r i) = As i" + assumes rc: "relChain r As" and AsB: "(\i \ Field r. As i) \ B" and Br: "|B| finite (Field r)" and cr: "Card_order r" + shows "\ i \ Field r. As (succ r i) = As i" proof(rule ccontr, auto) have 0: "wo_rel r" and 00: "Well_order r" - unfolding wo_rel_def by (metis card_order_on_well_order_on cr)+ + unfolding wo_rel_def by (metis card_order_on_well_order_on cr)+ have L: "isLimOrd r" using ir cr by (metis card_order_infinite_isLimOrd) have AsBs: "(\i \ Field r. As (succ r i)) \ B" - using AsB L apply safe by (metis "0" UN_I subsetD wo_rel.isLimOrd_succ) + using AsB L by (simp add: "0" Sup_le_iff wo_rel.isLimOrd_succ) assume As_s: "\i\Field r. As (succ r i) \ As i" have 1: "\i j. (i,j) \ r \ i \ j \ As i \ As j" proof safe @@ -1576,18 +1398,19 @@ hence rij: "(succ r i, j) \ r" by (metis "0" wo_rel.succ_smallest) hence "As (succ r i) \ As j" using rc unfolding relChain_def by auto moreover - {have "(i,succ r i) \ r" apply(rule wo_rel.succ_in[OF 0]) - using 1 unfolding aboveS_def by auto - hence "As i \ As (succ r i)" using As_s rc rij unfolding relChain_def Field_def by auto + { have "(i,succ r i) \ r" + by (meson "0" "1"(1) FieldI1 L wo_rel.isLimOrd_aboveS wo_rel.succ_in) + hence "As i \ As (succ r i)" using As_s rc rij unfolding relChain_def Field_def by auto } ultimately show False unfolding Asij by auto qed (insert rc, unfold relChain_def, auto) hence "\ i \ Field r. \ a. a \ As (succ r i) - As i" - using wo_rel.succ_in[OF 0] AsB - by(metis 0 card_order_infinite_isLimOrd cr ir psubset_imp_ex_mem - wo_rel.isLimOrd_aboveS wo_rel.succ_diff) + using wo_rel.succ_in[OF 0] AsB + by(metis 0 card_order_infinite_isLimOrd cr ir psubset_imp_ex_mem + wo_rel.isLimOrd_aboveS wo_rel.succ_diff) then obtain f where f: "\ i \ Field r. f i \ As (succ r i) - As i" by metis - have "inj_on f (Field r)" unfolding inj_on_def proof safe + have "inj_on f (Field r)" unfolding inj_on_def + proof safe fix i j assume ij: "i \ Field r" "j \ Field r" and fij: "f i = f j" show "i = j" proof(cases rule: wo_rel.cases_Total3[OF 0], safe) @@ -1610,118 +1433,62 @@ subsection \Regular vs. stable cardinals\ -lemma stable_regularCard: -assumes cr: "Card_order r" and ir: "\finite (Field r)" and st: "stable r" -shows "regularCard r" -unfolding regularCard_def proof safe - fix K assume K: "K \ Field r" and cof: "cofinal K r" - have "|K| \o r" using K by (metis card_of_Field_ordIso card_of_mono1 cr ordLeq_ordIso_trans) - moreover - {assume Kr: "|K| a \ Field r. f a \ K \ a \ f a \ (a, f a) \ r" - using cof unfolding cofinal_def by metis - hence "Field r \ (\a \ K. underS r a)" unfolding underS_def by auto - hence "r \o |\a \ K. underS r a|" using cr - by (metis Card_order_iff_ordLeq_card_of card_of_mono1 ordLeq_transitive) - also have "|\a \ K. underS r a| \o |SIGMA a: K. underS r a|" by (rule card_of_UNION_Sigma) - also - {have "\ a \ K. |underS r a| 'a set" - assume "|A| a\A. |F a| (\a \ A. finite(F a))" - by (auto simp add: finite_iff_ordLess_natLeq) - hence "finite(Sigma A F)" by (simp only: finite_SigmaI[of A F]) - thus "|Sigma A F | finite (Field r)" -shows "stable(cardSuc r)" -using infinite_cardSuc_regularCard regularCard_stable -by (metis CARD INF cardSuc_Card_order cardSuc_finite) - -lemma stable_ordIso1: -assumes ST: "stable r" and ISO: "r' =o r" -shows "stable r'" -proof(unfold stable_def, auto) - fix A::"'b set" and F::"'b \ 'b set" - assume "|A| a \ A. |F a| (\a \ A. |F a| finite (Field r)" + shows "stable(cardSuc r)" + using infinite_cardSuc_regularCard regularCard_stable + by (metis CARD INF cardSuc_Card_order cardSuc_finite) lemma stable_ordIso: -assumes "r =o r'" -shows "stable r = stable r'" -using assms stable_ordIso1 stable_ordIso2 by blast + assumes "r =o r'" + shows "stable r = stable r'" + by (metis assms ordIso_symmetric stable_ordIso1) lemma stable_nat: "stable |UNIV::nat set|" -using stable_natLeq card_of_nat stable_ordIso by auto + using stable_natLeq card_of_nat stable_ordIso by auto text\Below, the type of "A" is not important -- we just had to choose an appropriate type to make "A" possible. What is important is that arbitrarily large infinite sets of stable cardinality exist.\ lemma infinite_stable_exists: -assumes CARD: "\r \ R. Card_order (r::'a rel)" -shows "\(A :: (nat + 'a set)set). + assumes CARD: "\r \ R. Card_order (r::'a rel)" + shows "\(A :: (nat + 'a set)set). \finite A \ stable |A| \ (\r \ R. r (A :: (nat + 'a set)set). \finite A \ stable |A| \ |UNIV::'a set| finite(?B <+> {})" using finite_Plus_iff by blast moreover have "stable |?B|" using stable_natLeq card_of_nat stable_ordIso1 by blast hence "stable |?B <+> {}|" using stable_ordIso card_of_Plus_empty1 by blast moreover - have "\finite(Field |?B| ) \ finite(Field |UNIV::'a set| )" using Case1 by simp + have "\finite(Field |?B| ) \ finite(Field |UNIV::'a set| )" using True by simp hence "|UNIV::'a set| {}|" using card_of_Plus_empty1 ordLess_ordIso_trans by blast ultimately show ?thesis by blast next - assume Case1: "\finite (UNIV::'a set)" + case False hence *: "\finite(Field |UNIV::'a set| )" by simp let ?B = "Field(cardSuc |UNIV::'a set| )" have 0: "|?B| =o |{} <+> ?B|" using card_of_Plus_empty2 by blast - have 1: "\finite ?B" using Case1 card_of_cardSuc_finite by blast + have 1: "\finite ?B" using False card_of_cardSuc_finite by blast hence 2: "\finite({} <+> ?B)" using 0 card_of_ordIso_finite by blast have "|?B| =o cardSuc |UNIV::'a set|" - using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast + using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast moreover have "stable(cardSuc |UNIV::'a set| )" - using stable_cardSuc * card_of_Card_order by blast + using stable_cardSuc * card_of_Card_order by blast ultimately have "stable |?B|" using stable_ordIso by blast hence 3: "stable |{} <+> ?B|" using stable_ordIso 0 by blast have "|UNIV::'a set| ?B|" using 0 ordLess_ordIso_trans by blast thus ?thesis using 2 3 by blast qed @@ -1729,9 +1496,9 @@ qed corollary infinite_regularCard_exists: -assumes CARD: "\r \ R. Card_order (r::'a rel)" -shows "\(A :: (nat + 'a set)set). + assumes CARD: "\r \ R. Card_order (r::'a rel)" + shows "\(A :: (nat + 'a set)set). \finite A \ regularCard |A| \ (\r \ R. r More on Injections, Bijections and Inverses\ theory Fun_More -imports Main + imports Main begin subsection \Purely functional properties\ (* unused *) -(*1*)lemma notIn_Un_bij_betw2: -assumes NIN: "b \ A" and NIN': "b' \ A'" and - BIJ: "bij_betw f A A'" -shows "bij_betw f (A \ {b}) (A' \ {b'}) = (f b = b')" -proof - assume "f b = b'" - thus "bij_betw f (A \ {b}) (A' \ {b'})" - using assms notIn_Un_bij_betw[of b A f A'] by auto -next - assume *: "bij_betw f (A \ {b}) (A' \ {b'})" - hence "f b \ A' \ {b'}" - unfolding bij_betw_def by auto - moreover - {assume "f b \ A'" - then obtain b1 where 1: "b1 \ A" and 2: "f b1 = f b" using BIJ - by (auto simp add: bij_betw_def) - hence "b = b1" using * - by (auto simp add: bij_betw_def inj_on_def) - with 1 NIN have False by auto - } - ultimately show "f b = b'" by blast -qed - -(* unused *) (*1*)lemma bij_betw_diff_singl: -assumes BIJ: "bij_betw f A A'" and IN: "a \ A" -shows "bij_betw f (A - {a}) (A' - {f a})" + assumes BIJ: "bij_betw f A A'" and IN: "a \ A" + shows "bij_betw f (A - {a}) (A' - {f a})" proof- let ?B = "A - {a}" let ?B' = "A' - {f a}" have "f a \ A'" using IN BIJ unfolding bij_betw_def by blast hence "a \ ?B \ f a \ ?B' \ A = ?B \ {a} \ A' = ?B' \ {f a}" - using IN by blast + using IN by blast thus ?thesis using notIn_Un_bij_betw3[of a ?B f ?B'] BIJ by simp qed @@ -54,120 +30,62 @@ (* unused *) (*1*)lemma bij_betw_inv_into_RIGHT: -assumes BIJ: "bij_betw f A A'" and SUB: "B' \ A'" -shows "f `((inv_into A f)`B') = B'" -using assms -proof(auto simp add: bij_betw_inv_into_right) - let ?f' = "(inv_into A f)" - fix a' assume *: "a' \ B'" - hence "a' \ A'" using SUB by auto - hence "a' = f (?f' a')" - using BIJ by (auto simp add: bij_betw_inv_into_right) - thus "a' \ f ` (?f' ` B')" using * by blast -qed + assumes BIJ: "bij_betw f A A'" and SUB: "B' \ A'" + shows "f `((inv_into A f)`B') = B'" + by (metis BIJ SUB bij_betw_imp_surj_on image_inv_into_cancel) + (* unused *) (*1*)lemma bij_betw_inv_into_RIGHT_LEFT: -assumes BIJ: "bij_betw f A A'" and SUB: "B' \ A'" and - IM: "(inv_into A f) ` B' = B" -shows "f ` B = B'" -proof- - have "f`((inv_into A f)` B') = B'" - using assms bij_betw_inv_into_RIGHT[of f A A' B'] by auto - thus ?thesis using IM by auto -qed + assumes BIJ: "bij_betw f A A'" and SUB: "B' \ A'" and + IM: "(inv_into A f) ` B' = B" + shows "f ` B = B'" + by (metis BIJ IM SUB bij_betw_inv_into_RIGHT) (* unused *) (*2*)lemma bij_betw_inv_into_twice: -assumes "bij_betw f A A'" -shows "\a \ A. inv_into A' (inv_into A f) a = f a" -proof - let ?f' = "inv_into A f" let ?f'' = "inv_into A' ?f'" - have 1: "bij_betw ?f' A' A" using assms - by (auto simp add: bij_betw_inv_into) - fix a assume *: "a \ A" - then obtain a' where 2: "a' \ A'" and 3: "?f' a' = a" - using 1 unfolding bij_betw_def by force - hence "?f'' a = a'" - using * 1 3 by (auto simp add: bij_betw_inv_into_left) - moreover have "f a = a'" using assms 2 3 - by (auto simp add: bij_betw_inv_into_right) - ultimately show "?f'' a = f a" by simp -qed + assumes "bij_betw f A A'" + shows "\a \ A. inv_into A' (inv_into A f) a = f a" + by (simp add: assms inv_into_inv_into_eq) subsection \Properties involving Hilbert choice\ (*1*)lemma bij_betw_inv_into_LEFT: -assumes BIJ: "bij_betw f A A'" and SUB: "B \ A" -shows "(inv_into A f)`(f ` B) = B" -using assms unfolding bij_betw_def using inv_into_image_cancel by force + assumes BIJ: "bij_betw f A A'" and SUB: "B \ A" + shows "(inv_into A f)`(f ` B) = B" + using assms unfolding bij_betw_def using inv_into_image_cancel by force (*1*)lemma bij_betw_inv_into_LEFT_RIGHT: -assumes BIJ: "bij_betw f A A'" and SUB: "B \ A" and - IM: "f ` B = B'" -shows "(inv_into A f) ` B' = B" -using assms bij_betw_inv_into_LEFT[of f A A' B] by fast + assumes BIJ: "bij_betw f A A'" and SUB: "B \ A" and + IM: "f ` B = B'" + shows "(inv_into A f) ` B' = B" + using assms bij_betw_inv_into_LEFT[of f A A' B] by fast subsection \Other facts\ (*3*)lemma atLeastLessThan_injective: -assumes "{0 ..< m::nat} = {0 ..< n}" -shows "m = n" -proof- - {assume "m < n" - hence "m \ {0 ..< n}" by auto - hence "{0 ..< m} < {0 ..< n}" by auto - hence False using assms by blast - } - moreover - {assume "n < m" - hence "n \ {0 ..< m}" by auto - hence "{0 ..< n} < {0 ..< m}" by auto - hence False using assms by blast - } - ultimately show ?thesis by force -qed + assumes "{0 ..< m::nat} = {0 ..< n}" + shows "m = n" + using assms atLeast0LessThan by force (*2*)lemma atLeastLessThan_injective2: -"bij_betw f {0 ..< m::nat} {0 ..< n} \ m = n" -using finite_atLeastLessThan[of m] finite_atLeastLessThan[of n] - card_atLeastLessThan[of m] card_atLeastLessThan[of n] - bij_betw_iff_card[of "{0 ..< m}" "{0 ..< n}"] by auto + "bij_betw f {0 ..< m::nat} {0 ..< n} \ m = n" + using bij_betw_same_card by fastforce (*2*)lemma atLeastLessThan_less_eq: -"({0.. {0.. n)" -unfolding ivl_subset by arith + "({0.. {0.. n)" + by auto (*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 fastforce - -(* unused *) -(*2*)lemma atLeastLessThan_less_eq3: -"(\f. inj_on f {0..<(m::nat)} \ f ` {0.. {0.. n)" -using atLeastLessThan_less_eq2 -proof(auto) - assume "m \ n" - hence "inj_on id {0.. id ` {0.. {0..f. inj_on f {0.. f ` {0.. {0.. {0.. n" + by (metis assms card_inj_on_le card_lessThan finite_lessThan lessThan_atLeast0) (* unused *) (*3*)lemma atLeastLessThan_less: -"({0.. {0.. {0.. = (m \ n \ m ~= n)" - using atLeastLessThan_less_eq atLeastLessThan_injective by blast - also have "\ = (m < n)" by auto - finally show ?thesis . -qed + "({0..Basics on Order-Like Relations\ theory Order_Relation_More -imports Main + imports Main begin subsection \The upper and lower bounds operators\ lemma aboveS_subset_above: "aboveS r a \ above r a" -by(auto simp add: aboveS_def above_def) + by(auto simp add: aboveS_def above_def) lemma AboveS_subset_Above: "AboveS r A \ Above r A" -by(auto simp add: AboveS_def Above_def) + by(auto simp add: AboveS_def Above_def) lemma UnderS_disjoint: "A Int (UnderS r A) = {}" -by(auto simp add: UnderS_def) + by(auto simp add: UnderS_def) lemma aboveS_notIn: "a \ aboveS r a" -by(auto simp add: aboveS_def) + by(auto simp add: aboveS_def) lemma Refl_above_in: "\Refl r; a \ Field r\ \ a \ above r a" -by(auto simp add: refl_on_def above_def) + by(auto simp add: refl_on_def above_def) lemma in_Above_under: "a \ Field r \ a \ Above r (under r a)" -by(auto simp add: Above_def under_def) + by(auto simp add: Above_def under_def) lemma in_Under_above: "a \ Field r \ a \ Under r (above r a)" -by(auto simp add: Under_def above_def) + by(auto simp add: Under_def above_def) lemma in_UnderS_aboveS: "a \ Field r \ a \ UnderS r (aboveS r a)" -by(auto simp add: UnderS_def aboveS_def) + by(auto simp add: UnderS_def aboveS_def) lemma UnderS_subset_Under: "UnderS r A \ Under r A" -by(auto simp add: UnderS_def Under_def) + by(auto simp add: UnderS_def Under_def) lemma subset_Above_Under: "B \ Field r \ B \ Above r (Under r B)" -by(auto simp add: Above_def Under_def) + by(auto simp add: Above_def Under_def) lemma subset_Under_Above: "B \ Field r \ B \ Under r (Above r B)" -by(auto simp add: Under_def Above_def) + by(auto simp add: Under_def Above_def) lemma subset_AboveS_UnderS: "B \ Field r \ B \ AboveS r (UnderS r B)" -by(auto simp add: AboveS_def UnderS_def) + by(auto simp add: AboveS_def UnderS_def) lemma subset_UnderS_AboveS: "B \ Field r \ B \ UnderS r (AboveS r B)" -by(auto simp add: UnderS_def AboveS_def) + by(auto simp add: UnderS_def AboveS_def) lemma Under_Above_Galois: -"\B \ Field r; C \ Field r\ \ (B \ Above r C) = (C \ Under r B)" -by(unfold Above_def Under_def, blast) + "\B \ Field r; C \ Field r\ \ (B \ Above r C) = (C \ Under r B)" + by(unfold Above_def Under_def, blast) lemma UnderS_AboveS_Galois: -"\B \ Field r; C \ Field r\ \ (B \ AboveS r C) = (C \ UnderS r B)" -by(unfold AboveS_def UnderS_def, blast) + "\B \ Field r; C \ Field r\ \ (B \ AboveS r C) = (C \ UnderS r B)" + by(unfold AboveS_def UnderS_def, blast) lemma Refl_above_aboveS: assumes REFL: "Refl r" and IN: "a \ Field r" @@ -79,20 +79,20 @@ with LIN IN order_on_defs[of "Field r" r] total_on_def[of "Field r" r] have "(a,b) \ r \ a = b" by blast thus "(a,b) \ r" - using LIN IN order_on_defs[of _ r] refl_on_def[of _ r] by auto + using LIN IN order_on_defs[of _ r] refl_on_def[of _ r] by auto next fix b assume "(b, a) \ r" thus "b \ Field r" - using LIN order_on_defs[of _ r] refl_on_def[of _ r] by blast + using LIN order_on_defs[of _ r] refl_on_def[of _ r] by blast next fix b assume "b \ a" "(a, b) \ r" thus "b \ Field r" - using LIN order_on_defs[of "Field r" r] refl_on_def[of "Field r" r] by blast + using LIN order_on_defs[of "Field r" r] refl_on_def[of "Field r" r] by blast qed lemma Linear_order_underS_above_Field: -assumes LIN: "Linear_order r" and IN: "a \ Field r" -shows "Field r = underS r a \ above r a" + assumes LIN: "Linear_order r" and IN: "a \ Field r" + shows "Field r = underS r a \ above r a" proof(unfold underS_def above_def, auto) assume "a \ Field r" "(a, a) \ r" with LIN IN order_on_defs[of _ r] refl_on_def[of _ r] @@ -102,174 +102,174 @@ with LIN IN order_on_defs[of "Field r" r] total_on_def[of "Field r" r] have "(b,a) \ r \ b = a" by blast thus "(b,a) \ r" - using LIN IN order_on_defs[of _ r] refl_on_def[of _ r] by auto + using LIN IN order_on_defs[of _ r] refl_on_def[of _ r] by auto next fix b assume "b \ a" "(b, a) \ r" thus "b \ Field r" - using LIN order_on_defs[of _ r] refl_on_def[of _ r] by blast + using LIN order_on_defs[of _ r] refl_on_def[of _ r] by blast next fix b assume "(a, b) \ r" thus "b \ Field r" - using LIN order_on_defs[of "Field r" r] refl_on_def[of "Field r" r] by blast + using LIN order_on_defs[of "Field r" r] refl_on_def[of "Field r" r] by blast qed lemma under_empty: "a \ Field r \ under r a = {}" -unfolding Field_def under_def by auto + unfolding Field_def under_def by auto lemma Under_Field: "Under r A \ Field r" -by(unfold Under_def Field_def, auto) + by(unfold Under_def Field_def, auto) lemma UnderS_Field: "UnderS r A \ Field r" -by(unfold UnderS_def Field_def, auto) + by(unfold UnderS_def Field_def, auto) lemma above_Field: "above r a \ Field r" -by(unfold above_def Field_def, auto) + by(unfold above_def Field_def, auto) lemma aboveS_Field: "aboveS r a \ Field r" -by(unfold aboveS_def Field_def, auto) + by(unfold aboveS_def Field_def, auto) lemma Above_Field: "Above r A \ Field r" -by(unfold Above_def Field_def, auto) + by(unfold Above_def Field_def, auto) lemma Refl_under_Under: -assumes REFL: "Refl r" and NE: "A \ {}" -shows "Under r A = (\a \ A. under r a)" + assumes REFL: "Refl r" and NE: "A \ {}" + shows "Under r A = (\a \ A. under r a)" proof show "Under r A \ (\a \ A. under r a)" - by(unfold Under_def under_def, auto) + by(unfold Under_def under_def, auto) next show "(\a \ A. under r a) \ Under r A" proof(auto) fix x assume *: "\xa \ A. x \ under r xa" hence "\xa \ A. (x,xa) \ r" - by (simp add: under_def) + by (simp add: under_def) moreover {from NE obtain a where "a \ A" by blast - with * have "x \ under r a" by simp - hence "x \ Field r" - using under_Field[of r a] by auto + with * have "x \ under r a" by simp + hence "x \ Field r" + using under_Field[of r a] by auto } ultimately show "x \ Under r A" - unfolding Under_def by auto + unfolding Under_def by auto qed qed lemma Refl_underS_UnderS: -assumes REFL: "Refl r" and NE: "A \ {}" -shows "UnderS r A = (\a \ A. underS r a)" + assumes REFL: "Refl r" and NE: "A \ {}" + shows "UnderS r A = (\a \ A. underS r a)" proof show "UnderS r A \ (\a \ A. underS r a)" - by(unfold UnderS_def underS_def, auto) + by(unfold UnderS_def underS_def, auto) next show "(\a \ A. underS r a) \ UnderS r A" proof(auto) fix x assume *: "\xa \ A. x \ underS r xa" hence "\xa \ A. x \ xa \ (x,xa) \ r" - by (auto simp add: underS_def) + by (auto simp add: underS_def) moreover {from NE obtain a where "a \ A" by blast - with * have "x \ underS r a" by simp - hence "x \ Field r" - using underS_Field[of _ r a] by auto + with * have "x \ underS r a" by simp + hence "x \ Field r" + using underS_Field[of _ r a] by auto } ultimately show "x \ UnderS r A" - unfolding UnderS_def by auto + unfolding UnderS_def by auto qed qed lemma Refl_above_Above: -assumes REFL: "Refl r" and NE: "A \ {}" -shows "Above r A = (\a \ A. above r a)" + assumes REFL: "Refl r" and NE: "A \ {}" + shows "Above r A = (\a \ A. above r a)" proof show "Above r A \ (\a \ A. above r a)" - by(unfold Above_def above_def, auto) + by(unfold Above_def above_def, auto) next show "(\a \ A. above r a) \ Above r A" proof(auto) fix x assume *: "\xa \ A. x \ above r xa" hence "\xa \ A. (xa,x) \ r" - by (simp add: above_def) + by (simp add: above_def) moreover {from NE obtain a where "a \ A" by blast - with * have "x \ above r a" by simp - hence "x \ Field r" - using above_Field[of r a] by auto + with * have "x \ above r a" by simp + hence "x \ Field r" + using above_Field[of r a] by auto } ultimately show "x \ Above r A" - unfolding Above_def by auto + unfolding Above_def by auto qed qed lemma Refl_aboveS_AboveS: -assumes REFL: "Refl r" and NE: "A \ {}" -shows "AboveS r A = (\a \ A. aboveS r a)" + assumes REFL: "Refl r" and NE: "A \ {}" + shows "AboveS r A = (\a \ A. aboveS r a)" proof show "AboveS r A \ (\a \ A. aboveS r a)" - by(unfold AboveS_def aboveS_def, auto) + by(unfold AboveS_def aboveS_def, auto) next show "(\a \ A. aboveS r a) \ AboveS r A" proof(auto) fix x assume *: "\xa \ A. x \ aboveS r xa" hence "\xa \ A. xa \ x \ (xa,x) \ r" - by (auto simp add: aboveS_def) + by (auto simp add: aboveS_def) moreover {from NE obtain a where "a \ A" by blast - with * have "x \ aboveS r a" by simp - hence "x \ Field r" - using aboveS_Field[of r a] by auto + with * have "x \ aboveS r a" by simp + hence "x \ Field r" + using aboveS_Field[of r a] by auto } ultimately show "x \ AboveS r A" - unfolding AboveS_def by auto + unfolding AboveS_def by auto qed qed lemma under_Under_singl: "under r a = Under r {a}" -by(unfold Under_def under_def, auto simp add: Field_def) + by(unfold Under_def under_def, auto simp add: Field_def) lemma underS_UnderS_singl: "underS r a = UnderS r {a}" -by(unfold UnderS_def underS_def, auto simp add: Field_def) + by(unfold UnderS_def underS_def, auto simp add: Field_def) lemma above_Above_singl: "above r a = Above r {a}" -by(unfold Above_def above_def, auto simp add: Field_def) + by(unfold Above_def above_def, auto simp add: Field_def) lemma aboveS_AboveS_singl: "aboveS r a = AboveS r {a}" -by(unfold AboveS_def aboveS_def, auto simp add: Field_def) + by(unfold AboveS_def aboveS_def, auto simp add: Field_def) lemma Under_decr: "A \ B \ Under r B \ Under r A" -by(unfold Under_def, auto) + by(unfold Under_def, auto) lemma UnderS_decr: "A \ B \ UnderS r B \ UnderS r A" -by(unfold UnderS_def, auto) + by(unfold UnderS_def, auto) lemma Above_decr: "A \ B \ Above r B \ Above r A" -by(unfold Above_def, auto) + by(unfold Above_def, auto) lemma AboveS_decr: "A \ B \ AboveS r B \ AboveS r A" -by(unfold AboveS_def, auto) + by(unfold AboveS_def, auto) lemma under_incl_iff: -assumes TRANS: "trans r" and REFL: "Refl r" and IN: "a \ Field r" -shows "(under r a \ under r b) = ((a,b) \ r)" + assumes TRANS: "trans r" and REFL: "Refl r" and IN: "a \ Field r" + shows "(under r a \ under r b) = ((a,b) \ r)" proof assume "(a,b) \ r" thus "under r a \ under r b" using TRANS - by (auto simp add: under_incr) + by (auto simp add: under_incr) next assume "under r a \ under r b" moreover have "a \ under r a" using REFL IN - by (auto simp add: Refl_under_in) + by (auto simp add: Refl_under_in) ultimately show "(a,b) \ r" - by (auto simp add: under_def) + by (auto simp add: under_def) qed lemma above_decr: -assumes TRANS: "trans r" and REL: "(a,b) \ r" -shows "above r b \ above r a" + assumes TRANS: "trans r" and REL: "(a,b) \ r" + shows "above r b \ above r a" proof(unfold above_def, auto) fix x assume "(b,x) \ r" with REL TRANS trans_def[of r] @@ -277,9 +277,9 @@ qed lemma aboveS_decr: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - REL: "(a,b) \ r" -shows "aboveS r b \ aboveS r a" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + REL: "(a,b) \ r" + shows "aboveS r b \ aboveS r a" proof(unfold aboveS_def, auto) assume *: "a \ b" and **: "(b,a) \ r" with ANTISYM antisym_def[of r] REL @@ -291,26 +291,26 @@ qed lemma under_trans: -assumes TRANS: "trans r" and - IN1: "a \ under r b" and IN2: "b \ under r c" -shows "a \ under r c" + assumes TRANS: "trans r" and + IN1: "a \ under r b" and IN2: "b \ under r c" + shows "a \ under r c" proof- have "(a,b) \ r \ (b,c) \ r" - using IN1 IN2 under_def by fastforce + using IN1 IN2 under_def by fastforce hence "(a,c) \ r" - using TRANS trans_def[of r] by blast + using TRANS trans_def[of r] by blast thus ?thesis unfolding under_def by simp qed lemma under_underS_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ under r b" and IN2: "b \ underS r c" -shows "a \ underS r c" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ under r b" and IN2: "b \ underS r c" + shows "a \ underS r c" proof- have 0: "(a,b) \ r \ (b,c) \ r" - using IN1 IN2 under_def underS_def by fastforce + using IN1 IN2 under_def underS_def by fastforce hence 1: "(a,c) \ r" - using TRANS trans_def[of r] by blast + using TRANS trans_def[of r] by blast have 2: "b \ c" using IN2 underS_def by force have 3: "a \ c" proof @@ -321,14 +321,14 @@ qed lemma underS_under_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ underS r b" and IN2: "b \ under r c" -shows "a \ underS r c" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ underS r b" and IN2: "b \ under r c" + shows "a \ underS r c" proof- have 0: "(a,b) \ r \ (b,c) \ r" - using IN1 IN2 under_def underS_def by fast + using IN1 IN2 under_def underS_def by fast hence 1: "(a,c) \ r" - using TRANS trans_def[of r] by fast + using TRANS trans_def[of r] by fast have 2: "a \ b" using IN1 underS_def by force have 3: "a \ c" proof @@ -339,36 +339,36 @@ qed lemma underS_underS_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ underS r b" and IN2: "b \ underS r c" -shows "a \ underS r c" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ underS r b" and IN2: "b \ underS r c" + shows "a \ underS r c" proof- have "a \ under r b" - using IN1 underS_subset_under by fast + using IN1 underS_subset_under by fast with assms under_underS_trans show ?thesis by fast qed lemma above_trans: -assumes TRANS: "trans r" and - IN1: "b \ above r a" and IN2: "c \ above r b" -shows "c \ above r a" + assumes TRANS: "trans r" and + IN1: "b \ above r a" and IN2: "c \ above r b" + shows "c \ above r a" proof- have "(a,b) \ r \ (b,c) \ r" - using IN1 IN2 above_def by fast + using IN1 IN2 above_def by fast hence "(a,c) \ r" - using TRANS trans_def[of r] by blast + using TRANS trans_def[of r] by blast thus ?thesis unfolding above_def by simp qed lemma above_aboveS_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "b \ above r a" and IN2: "c \ aboveS r b" -shows "c \ aboveS r a" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "b \ above r a" and IN2: "c \ aboveS r b" + shows "c \ aboveS r a" proof- have 0: "(a,b) \ r \ (b,c) \ r" - using IN1 IN2 above_def aboveS_def by fast + using IN1 IN2 above_def aboveS_def by fast hence 1: "(a,c) \ r" - using TRANS trans_def[of r] by blast + using TRANS trans_def[of r] by blast have 2: "b \ c" using IN2 aboveS_def by force have 3: "a \ c" proof @@ -379,14 +379,14 @@ qed lemma aboveS_above_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "b \ aboveS r a" and IN2: "c \ above r b" -shows "c \ aboveS r a" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "b \ aboveS r a" and IN2: "c \ above r b" + shows "c \ aboveS r a" proof- have 0: "(a,b) \ r \ (b,c) \ r" - using IN1 IN2 above_def aboveS_def by fast + using IN1 IN2 above_def aboveS_def by fast hence 1: "(a,c) \ r" - using TRANS trans_def[of r] by blast + using TRANS trans_def[of r] by blast have 2: "a \ b" using IN1 aboveS_def by force have 3: "a \ c" proof @@ -397,26 +397,26 @@ qed lemma aboveS_aboveS_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "b \ aboveS r a" and IN2: "c \ aboveS r b" -shows "c \ aboveS r a" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "b \ aboveS r a" and IN2: "c \ aboveS r b" + shows "c \ aboveS r a" proof- have "b \ above r a" - using IN1 aboveS_subset_above by fast + using IN1 aboveS_subset_above by fast with assms above_aboveS_trans show ?thesis by fast qed lemma under_Under_trans: -assumes TRANS: "trans r" and - IN1: "a \ under r b" and IN2: "b \ Under r C" -shows "a \ Under r C" + assumes TRANS: "trans r" and + IN1: "a \ under r b" and IN2: "b \ Under r C" + shows "a \ Under r C" proof- have "b \ {u \ Field r. \x. x \ C \ (u, x) \ r}" using IN2 Under_def by force hence "(a,b) \ r \ (\c \ C. (b,c) \ r)" using IN1 under_def by fast hence "\c \ C. (a,c) \ r" - using TRANS trans_def[of r] by blast + using TRANS trans_def[of r] by blast moreover have "a \ Field r" using IN1 unfolding Field_def under_def by blast ultimately @@ -424,53 +424,53 @@ qed lemma underS_Under_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ underS r b" and IN2: "b \ Under r C" -shows "a \ UnderS r C" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ underS r b" and IN2: "b \ Under r C" + shows "a \ UnderS r C" proof- from IN1 have "a \ under r b" - using underS_subset_under[of r b] by fast + using underS_subset_under[of r b] by fast with assms under_Under_trans have "a \ Under r C" by fast - (* *) + (* *) moreover have "a \ C" proof assume *: "a \ C" have 1: "b \ a \ (a,b) \ r" - using IN1 underS_def[of r b] by auto + using IN1 underS_def[of r b] by auto have "\c \ C. (b,c) \ r" - using IN2 Under_def[of r C] by auto + using IN2 Under_def[of r C] by auto with * have "(b,a) \ r" by simp with 1 ANTISYM antisym_def[of r] show False by blast qed - (* *) + (* *) ultimately show ?thesis unfolding UnderS_def - using Under_def by force + using Under_def by force qed lemma underS_UnderS_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ underS r b" and IN2: "b \ UnderS r C" -shows "a \ UnderS r C" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ underS r b" and IN2: "b \ UnderS r C" + shows "a \ UnderS r C" proof- from IN2 have "b \ Under r C" - using UnderS_subset_Under[of r C] by blast + using UnderS_subset_Under[of r C] by blast with underS_Under_trans assms show ?thesis by force qed lemma above_Above_trans: -assumes TRANS: "trans r" and - IN1: "a \ above r b" and IN2: "b \ Above r C" -shows "a \ Above r C" + assumes TRANS: "trans r" and + IN1: "a \ above r b" and IN2: "b \ Above r C" + shows "a \ Above r C" proof- have "(b,a) \ r \ (\c \ C. (c,b) \ r)" using IN1[unfolded above_def] IN2[unfolded Above_def] by simp hence "\c \ C. (c,a) \ r" - using TRANS trans_def[of r] by blast + using TRANS trans_def[of r] by blast moreover have "a \ Field r" using IN1[unfolded above_def] Field_def by fast ultimately @@ -478,95 +478,95 @@ qed lemma aboveS_Above_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ aboveS r b" and IN2: "b \ Above r C" -shows "a \ AboveS r C" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ aboveS r b" and IN2: "b \ Above r C" + shows "a \ AboveS r C" proof- from IN1 have "a \ above r b" - using aboveS_subset_above[of r b] by blast + using aboveS_subset_above[of r b] by blast with assms above_Above_trans have "a \ Above r C" by fast - (* *) + (* *) moreover have "a \ C" proof assume *: "a \ C" have 1: "b \ a \ (b,a) \ r" - using IN1 aboveS_def[of r b] by auto + using IN1 aboveS_def[of r b] by auto have "\c \ C. (c,b) \ r" - using IN2 Above_def[of r C] by auto + using IN2 Above_def[of r C] by auto with * have "(a,b) \ r" by simp with 1 ANTISYM antisym_def[of r] show False by blast qed - (* *) + (* *) ultimately show ?thesis unfolding AboveS_def - using Above_def by force + using Above_def by force qed lemma above_AboveS_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ above r b" and IN2: "b \ AboveS r C" -shows "a \ AboveS r C" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ above r b" and IN2: "b \ AboveS r C" + shows "a \ AboveS r C" proof- from IN2 have "b \ Above r C" - using AboveS_subset_Above[of r C] by blast + using AboveS_subset_Above[of r C] by blast with assms above_Above_trans have "a \ Above r C" by force - (* *) + (* *) moreover have "a \ C" proof assume *: "a \ C" have 1: "(b,a) \ r" - using IN1 above_def[of r b] by auto + using IN1 above_def[of r b] by auto have "\c \ C. b \ c \ (c,b) \ r" - using IN2 AboveS_def[of r C] by auto + using IN2 AboveS_def[of r C] by auto with * have "b \ a \ (a,b) \ r" by simp with 1 ANTISYM antisym_def[of r] show False by blast qed - (* *) + (* *) ultimately show ?thesis unfolding AboveS_def - using Above_def by force + using Above_def by force qed lemma aboveS_AboveS_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ aboveS r b" and IN2: "b \ AboveS r C" -shows "a \ AboveS r C" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ aboveS r b" and IN2: "b \ AboveS r C" + shows "a \ AboveS r C" proof- from IN2 have "b \ Above r C" - using AboveS_subset_Above[of r C] by blast + using AboveS_subset_Above[of r C] by blast with aboveS_Above_trans assms show ?thesis by force qed lemma under_UnderS_trans: -assumes TRANS: "trans r" and ANTISYM: "antisym r" and - IN1: "a \ under r b" and IN2: "b \ UnderS r C" -shows "a \ UnderS r C" + assumes TRANS: "trans r" and ANTISYM: "antisym r" and + IN1: "a \ under r b" and IN2: "b \ UnderS r C" + shows "a \ UnderS r C" proof- from IN2 have "b \ Under r C" - using UnderS_subset_Under[of r C] by blast + using UnderS_subset_Under[of r C] by blast with assms under_Under_trans have "a \ Under r C" by force - (* *) + (* *) moreover have "a \ C" proof assume *: "a \ C" have 1: "(a,b) \ r" - using IN1 under_def[of r b] by auto + using IN1 under_def[of r b] by auto have "\c \ C. b \ c \ (b,c) \ r" - using IN2 UnderS_def[of r C] by blast + using IN2 UnderS_def[of r C] by blast with * have "b \ a \ (b,a) \ r" by blast with 1 ANTISYM antisym_def[of r] show False by blast qed - (* *) + (* *) ultimately show ?thesis unfolding UnderS_def Under_def by fast qed @@ -575,12 +575,12 @@ subsection \Properties depending on more than one relation\ lemma under_incr2: -"r \ r' \ under r a \ under r' a" -unfolding under_def by blast + "r \ r' \ under r a \ under r' a" + unfolding under_def by blast lemma underS_incr2: -"r \ r' \ underS r a \ underS r' a" -unfolding underS_def by blast + "r \ r' \ underS r a \ underS r' a" + unfolding underS_def by blast (* FIXME: r \ r' lemma Under_incr: @@ -601,12 +601,12 @@ *) lemma above_incr2: -"r \ r' \ above r a \ above r' a" -unfolding above_def by blast + "r \ r' \ above r a \ above r' a" + unfolding above_def by blast lemma aboveS_incr2: -"r \ r' \ aboveS r a \ aboveS r' a" -unfolding aboveS_def by blast + "r \ r' \ aboveS r a \ aboveS r' a" + unfolding aboveS_def by blast (* FIXME lemma Above_incr: diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Cardinals/Order_Union.thy --- a/src/HOL/Cardinals/Order_Union.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Cardinals/Order_Union.thy Sun Jan 15 20:00:44 2023 +0100 @@ -7,7 +7,7 @@ section \Order Union\ theory Order_Union -imports Main + imports Main begin definition Osum :: "'a rel \ 'a rel \ 'a rel" (infix "Osum" 60) where @@ -19,10 +19,10 @@ unfolding Osum_def Field_def by blast lemma Osum_wf: -assumes FLD: "Field r Int Field r' = {}" and - WF: "wf r" and WF': "wf r'" -shows "wf (r Osum r')" -unfolding wf_eq_minimal2 unfolding Field_Osum + assumes FLD: "Field r Int Field r' = {}" and + WF: "wf r" and WF': "wf r'" + shows "wf (r Osum r')" + unfolding wf_eq_minimal2 unfolding Field_Osum proof(intro allI impI, elim conjE) fix A assume *: "A \ Field r \ Field r'" and **: "A \ {}" obtain B where B_def: "B = A Int Field r" by blast @@ -31,26 +31,26 @@ assume Case1: "B \ {}" hence "B \ {} \ B \ Field r" using B_def by auto then obtain a where 1: "a \ B" and 2: "\a1 \ B. (a1,a) \ r" - using WF unfolding wf_eq_minimal2 by blast + using WF unfolding wf_eq_minimal2 by blast hence 3: "a \ Field r \ a \ Field r'" using B_def FLD by auto - (* *) + (* *) have "\a1 \ A. (a1,a) \ r Osum r'" proof(intro ballI) fix a1 assume **: "a1 \ A" {assume Case11: "a1 \ Field r" - hence "(a1,a) \ r" using B_def ** 2 by auto - moreover - have "(a1,a) \ r'" using 3 by (auto simp add: Field_def) - ultimately have "(a1,a) \ r Osum r'" - using 3 unfolding Osum_def by auto + hence "(a1,a) \ r" using B_def ** 2 by auto + moreover + have "(a1,a) \ r'" using 3 by (auto simp add: Field_def) + ultimately have "(a1,a) \ r Osum r'" + using 3 unfolding Osum_def by auto } moreover {assume Case12: "a1 \ Field r" - hence "(a1,a) \ r" unfolding Field_def by auto - moreover - have "(a1,a) \ r'" using 3 unfolding Field_def by auto - ultimately have "(a1,a) \ r Osum r'" - using 3 unfolding Osum_def by auto + hence "(a1,a) \ r" unfolding Field_def by auto + moreover + have "(a1,a) \ r'" using 3 unfolding Field_def by auto + ultimately have "(a1,a) \ r Osum r'" + using 3 unfolding Osum_def by auto } ultimately show "(a1,a) \ r Osum r'" by blast qed @@ -59,9 +59,9 @@ assume Case2: "B = {}" hence 1: "A \ {} \ A \ Field r'" using * ** B_def by auto then obtain a' where 2: "a' \ A" and 3: "\a1' \ A. (a1',a') \ r'" - using WF' unfolding wf_eq_minimal2 by blast + using WF' unfolding wf_eq_minimal2 by blast hence 4: "a' \ Field r' \ a' \ Field r" using 1 FLD by blast - (* *) + (* *) have "\a1' \ A. (a1',a') \ r Osum r'" proof(unfold Osum_def, auto simp add: 3) fix a1' assume "(a1', a') \ r" @@ -75,296 +75,118 @@ qed lemma Osum_Refl: -assumes FLD: "Field r Int Field r' = {}" and - REFL: "Refl r" and REFL': "Refl r'" -shows "Refl (r Osum r')" -using assms -unfolding refl_on_def Field_Osum unfolding Osum_def by blast + assumes FLD: "Field r Int Field r' = {}" and + REFL: "Refl r" and REFL': "Refl r'" + shows "Refl (r Osum r')" + using assms + unfolding refl_on_def Field_Osum unfolding Osum_def by blast lemma Osum_trans: -assumes FLD: "Field r Int Field r' = {}" and - TRANS: "trans r" and TRANS': "trans r'" -shows "trans (r Osum r')" -proof(unfold trans_def, auto) - fix x y z assume *: "(x, y) \ r \o r'" and **: "(y, z) \ r \o r'" - show "(x, z) \ r \o r'" - proof- - {assume Case1: "(x,y) \ r" - hence 1: "x \ Field r \ y \ Field r" unfolding Field_def by auto - have ?thesis - proof- - {assume Case11: "(y,z) \ r" - hence "(x,z) \ r" using Case1 TRANS trans_def[of r] by blast - hence ?thesis unfolding Osum_def by auto - } - moreover - {assume Case12: "(y,z) \ r'" - hence "y \ Field r'" unfolding Field_def by auto - hence False using FLD 1 by auto - } - moreover - {assume Case13: "z \ Field r'" - hence ?thesis using 1 unfolding Osum_def by auto - } - ultimately show ?thesis using ** unfolding Osum_def by blast - qed - } - moreover - {assume Case2: "(x,y) \ r'" - hence 2: "x \ Field r' \ y \ Field r'" unfolding Field_def by auto - have ?thesis - proof- - {assume Case21: "(y,z) \ r" - hence "y \ Field r" unfolding Field_def by auto - hence False using FLD 2 by auto - } - moreover - {assume Case22: "(y,z) \ r'" - hence "(x,z) \ r'" using Case2 TRANS' trans_def[of r'] by blast - hence ?thesis unfolding Osum_def by auto - } - moreover - {assume Case23: "y \ Field r" - hence False using FLD 2 by auto - } - ultimately show ?thesis using ** unfolding Osum_def by blast - qed - } - moreover - {assume Case3: "x \ Field r \ y \ Field r'" - have ?thesis - proof- - {assume Case31: "(y,z) \ r" - hence "y \ Field r" unfolding Field_def by auto - hence False using FLD Case3 by auto - } - moreover - {assume Case32: "(y,z) \ r'" - hence "z \ Field r'" unfolding Field_def by blast - hence ?thesis unfolding Osum_def using Case3 by auto - } - moreover - {assume Case33: "y \ Field r" - hence False using FLD Case3 by auto - } - ultimately show ?thesis using ** unfolding Osum_def by blast - qed - } - ultimately show ?thesis using * unfolding Osum_def by blast - qed -qed + assumes FLD: "Field r Int Field r' = {}" and + TRANS: "trans r" and TRANS': "trans r'" + shows "trans (r Osum r')" + using assms unfolding Osum_def trans_def disjoint_iff Field_iff by blast lemma Osum_Preorder: -"\Field r Int Field r' = {}; Preorder r; Preorder r'\ \ Preorder (r Osum r')" -unfolding preorder_on_def using Osum_Refl Osum_trans by blast + "\Field r Int Field r' = {}; Preorder r; Preorder r'\ \ Preorder (r Osum r')" + unfolding preorder_on_def using Osum_Refl Osum_trans by blast lemma Osum_antisym: -assumes FLD: "Field r Int Field r' = {}" and - AN: "antisym r" and AN': "antisym r'" -shows "antisym (r Osum r')" -proof(unfold antisym_def, auto) - fix x y assume *: "(x, y) \ r \o r'" and **: "(y, x) \ r \o r'" - show "x = y" - proof- - {assume Case1: "(x,y) \ r" - hence 1: "x \ Field r \ y \ Field r" unfolding Field_def by auto - have ?thesis - proof- - have "(y,x) \ r \ ?thesis" - using Case1 AN antisym_def[of r] by blast - moreover - {assume "(y,x) \ r'" - hence "y \ Field r'" unfolding Field_def by auto - hence False using FLD 1 by auto - } - moreover - have "x \ Field r' \ False" using FLD 1 by auto - ultimately show ?thesis using ** unfolding Osum_def by blast - qed - } - moreover - {assume Case2: "(x,y) \ r'" - hence 2: "x \ Field r' \ y \ Field r'" unfolding Field_def by auto - have ?thesis - proof- - {assume "(y,x) \ r" - hence "y \ Field r" unfolding Field_def by auto - hence False using FLD 2 by auto - } - moreover - have "(y,x) \ r' \ ?thesis" - using Case2 AN' antisym_def[of r'] by blast - moreover - {assume "y \ Field r" - hence False using FLD 2 by auto - } - ultimately show ?thesis using ** unfolding Osum_def by blast - qed - } - moreover - {assume Case3: "x \ Field r \ y \ Field r'" - have ?thesis - proof- - {assume "(y,x) \ r" - hence "y \ Field r" unfolding Field_def by auto - hence False using FLD Case3 by auto - } - moreover - {assume Case32: "(y,x) \ r'" - hence "x \ Field r'" unfolding Field_def by blast - hence False using FLD Case3 by auto - } - moreover - have "\ y \ Field r" using FLD Case3 by auto - ultimately show ?thesis using ** unfolding Osum_def by blast - qed - } - ultimately show ?thesis using * unfolding Osum_def by blast - qed -qed + assumes FLD: "Field r Int Field r' = {}" and + AN: "antisym r" and AN': "antisym r'" + shows "antisym (r Osum r')" + using assms by (auto simp: disjoint_iff antisym_def Osum_def Field_def) lemma Osum_Partial_order: -"\Field r Int Field r' = {}; Partial_order r; Partial_order r'\ \ + "\Field r Int Field r' = {}; Partial_order r; Partial_order r'\ \ Partial_order (r Osum r')" -unfolding partial_order_on_def using Osum_Preorder Osum_antisym by blast + unfolding partial_order_on_def using Osum_Preorder Osum_antisym by blast lemma Osum_Total: -assumes FLD: "Field r Int Field r' = {}" and - TOT: "Total r" and TOT': "Total r'" -shows "Total (r Osum r')" -using assms -unfolding total_on_def Field_Osum unfolding Osum_def by blast + assumes FLD: "Field r Int Field r' = {}" and + TOT: "Total r" and TOT': "Total r'" + shows "Total (r Osum r')" + using assms + unfolding total_on_def Field_Osum unfolding Osum_def by blast lemma Osum_Linear_order: -"\Field r Int Field r' = {}; Linear_order r; Linear_order r'\ \ - Linear_order (r Osum r')" -unfolding linear_order_on_def using Osum_Partial_order Osum_Total by blast + "\Field r Int Field r' = {}; Linear_order r; Linear_order r'\ \ Linear_order (r Osum r')" + by (simp add: Osum_Partial_order Osum_Total linear_order_on_def) lemma Osum_minus_Id1: -assumes "r \ Id" -shows "(r Osum r') - Id \ (r' - Id) \ (Field r \ Field r')" -proof- - let ?Left = "(r Osum r') - Id" - let ?Right = "(r' - Id) \ (Field r \ Field r')" - {fix a::'a and b assume *: "(a,b) \ Id" - {assume "(a,b) \ r" - with * have False using assms by auto - } - moreover - {assume "(a,b) \ r'" - with * have "(a,b) \ r' - Id" by auto - } - ultimately - have "(a,b) \ ?Left \ (a,b) \ ?Right" - unfolding Osum_def by auto - } - thus ?thesis by auto -qed + assumes "r \ Id" + shows "(r Osum r') - Id \ (r' - Id) \ (Field r \ Field r')" +using assms by (force simp: Osum_def) lemma Osum_minus_Id2: -assumes "r' \ Id" -shows "(r Osum r') - Id \ (r - Id) \ (Field r \ Field r')" -proof- - let ?Left = "(r Osum r') - Id" - let ?Right = "(r - Id) \ (Field r \ Field r')" - {fix a::'a and b assume *: "(a,b) \ Id" - {assume "(a,b) \ r'" - with * have False using assms by auto - } - moreover - {assume "(a,b) \ r" - with * have "(a,b) \ r - Id" by auto - } - ultimately - have "(a,b) \ ?Left \ (a,b) \ ?Right" - unfolding Osum_def by auto - } - thus ?thesis by auto -qed + assumes "r' \ Id" + shows "(r Osum r') - Id \ (r - Id) \ (Field r \ Field r')" +using assms by (force simp: Osum_def) lemma Osum_minus_Id: -assumes TOT: "Total r" and TOT': "Total r'" and - NID: "\ (r \ Id)" and NID': "\ (r' \ Id)" -shows "(r Osum r') - Id \ (r - Id) Osum (r' - Id)" -proof- - {fix a a' assume *: "(a,a') \ (r Osum r')" and **: "a \ a'" - have "(a,a') \ (r - Id) Osum (r' - Id)" - proof- - {assume "(a,a') \ r \ (a,a') \ r'" - with ** have ?thesis unfolding Osum_def by auto - } - moreover - {assume "a \ Field r \ a' \ Field r'" - hence "a \ Field(r - Id) \ a' \ Field (r' - Id)" - using assms Total_Id_Field by blast - hence ?thesis unfolding Osum_def by auto - } - ultimately show ?thesis using * unfolding Osum_def by fast - qed - } - thus ?thesis by(auto simp add: Osum_def) -qed + assumes TOT: "Total r" and TOT': "Total r'" and + NID: "\ (r \ Id)" and NID': "\ (r' \ Id)" + shows "(r Osum r') - Id \ (r - Id) Osum (r' - Id)" + using assms Total_Id_Field by (force simp: Osum_def) lemma wf_Int_Times: -assumes "A Int B = {}" -shows "wf(A \ B)" -unfolding wf_def using assms by blast + assumes "A Int B = {}" + shows "wf(A \ B)" + unfolding wf_def using assms by blast lemma Osum_wf_Id: -assumes TOT: "Total r" and TOT': "Total r'" and - FLD: "Field r Int Field r' = {}" and - WF: "wf(r - Id)" and WF': "wf(r' - Id)" -shows "wf ((r Osum r') - Id)" + assumes TOT: "Total r" and TOT': "Total r'" and + FLD: "Field r Int Field r' = {}" and + WF: "wf(r - Id)" and WF': "wf(r' - Id)" + shows "wf ((r Osum r') - Id)" proof(cases "r \ Id \ r' \ Id") assume Case1: "\(r \ Id \ r' \ Id)" have "Field(r - Id) Int Field(r' - Id) = {}" - using FLD mono_Field[of "r - Id" r] mono_Field[of "r' - Id" r'] - Diff_subset[of r Id] Diff_subset[of r' Id] by blast + using Case1 FLD TOT TOT' Total_Id_Field by blast thus ?thesis - using Case1 Osum_minus_Id[of r r'] assms Osum_wf[of "r - Id" "r' - Id"] - wf_subset[of "(r - Id) \o (r' - Id)" "(r Osum r') - Id"] by auto + by (meson Case1 Osum_minus_Id Osum_wf TOT TOT' WF WF' wf_subset) next have 1: "wf(Field r \ Field r')" - using FLD by (auto simp add: wf_Int_Times) + using FLD by (auto simp add: wf_Int_Times) assume Case2: "r \ Id \ r' \ Id" moreover {assume Case21: "r \ Id" - hence "(r Osum r') - Id \ (r' - Id) \ (Field r \ Field r')" - using Osum_minus_Id1[of r r'] by simp - moreover - {have "Domain(Field r \ Field r') Int Range(r' - Id) = {}" - using FLD unfolding Field_def by blast - hence "wf((r' - Id) \ (Field r \ Field r'))" - using 1 WF' wf_Un[of "Field r \ Field r'" "r' - Id"] - by (auto simp add: Un_commute) - } - ultimately have ?thesis using wf_subset by blast + hence "(r Osum r') - Id \ (r' - Id) \ (Field r \ Field r')" + using Osum_minus_Id1[of r r'] by simp + moreover + {have "Domain(Field r \ Field r') Int Range(r' - Id) = {}" + using FLD unfolding Field_def by blast + hence "wf((r' - Id) \ (Field r \ Field r'))" + using 1 WF' wf_Un[of "Field r \ Field r'" "r' - Id"] + by (auto simp add: Un_commute) + } + ultimately have ?thesis using wf_subset by blast } moreover {assume Case22: "r' \ Id" - hence "(r Osum r') - Id \ (r - Id) \ (Field r \ Field r')" - using Osum_minus_Id2[of r' r] by simp - moreover - {have "Range(Field r \ Field r') Int Domain(r - Id) = {}" - using FLD unfolding Field_def by blast - hence "wf((r - Id) \ (Field r \ Field r'))" - using 1 WF wf_Un[of "r - Id" "Field r \ Field r'"] - by (auto simp add: Un_commute) - } - ultimately have ?thesis using wf_subset by blast + hence "(r Osum r') - Id \ (r - Id) \ (Field r \ Field r')" + using Osum_minus_Id2[of r' r] by simp + moreover + {have "Range(Field r \ Field r') Int Domain(r - Id) = {}" + using FLD unfolding Field_def by blast + hence "wf((r - Id) \ (Field r \ Field r'))" + using 1 WF wf_Un[of "r - Id" "Field r \ Field r'"] + by (auto simp add: Un_commute) + } + ultimately have ?thesis using wf_subset by blast } ultimately show ?thesis by blast qed lemma Osum_Well_order: -assumes FLD: "Field r Int Field r' = {}" and - WELL: "Well_order r" and WELL': "Well_order r'" -shows "Well_order (r Osum r')" + assumes FLD: "Field r Int Field r' = {}" and + WELL: "Well_order r" and WELL': "Well_order r'" + shows "Well_order (r Osum r')" proof- have "Total r \ Total r'" using WELL WELL' - by (auto simp add: order_on_defs) + by (auto simp add: order_on_defs) thus ?thesis using assms unfolding well_order_on_def - using Osum_Linear_order Osum_wf_Id by blast + using Osum_Linear_order Osum_wf_Id by blast qed end diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Cardinals/Ordinal_Arithmetic.thy --- a/src/HOL/Cardinals/Ordinal_Arithmetic.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy Sun Jan 15 20:00:44 2023 +0100 @@ -8,12 +8,12 @@ section \Ordinal Arithmetic\ theory Ordinal_Arithmetic -imports Wellorder_Constructions + imports Wellorder_Constructions begin definition osum :: "'a rel \ 'b rel \ ('a + 'b) rel" (infixr "+o" 70) -where - "r +o r' = map_prod Inl Inl ` r \ map_prod Inr Inr ` r' \ + where + "r +o r' = map_prod Inl Inl ` r \ map_prod Inr Inr ` r' \ {(Inl a, Inr a') | a a' . a \ Field r \ a' \ Field r'}" lemma Field_osum: "Field(r +o r') = Inl ` Field r \ Inr ` Field r'" @@ -24,9 +24,10 @@ unfolding refl_on_def Field_osum unfolding osum_def by blast lemma osum_trans: -assumes TRANS: "trans r" and TRANS': "trans r'" -shows "trans (r +o r')" -proof(unfold trans_def, safe) + assumes TRANS: "trans r" and TRANS': "trans r'" + shows "trans (r +o r')" + unfolding trans_def +proof(safe) fix x y z assume *: "(x, y) \ r +o r'" "(y, z) \ r +o r'" thus "(x, z) \ r +o r'" proof (cases x y z rule: sum.exhaust[case_product sum.exhaust sum.exhaust]) @@ -66,9 +67,9 @@ unfolding linear_order_on_def using osum_Partial_order osum_Total by blast lemma osum_wf: -assumes WF: "wf r" and WF': "wf r'" -shows "wf (r +o r')" -unfolding wf_eq_minimal2 unfolding Field_osum + assumes WF: "wf r" and WF': "wf r'" + shows "wf (r +o r')" + unfolding wf_eq_minimal2 unfolding Field_osum proof(intro allI impI, elim conjE) fix A assume *: "A \ Inl ` Field r \ Inr ` Field r'" and **: "A \ {}" obtain B where B_def: "B = A Int Inl ` Field r" by blast @@ -113,8 +114,8 @@ proof(cases "r \ Id \ r' \ Id") case False thus ?thesis - using osum_minus_Id[of r r'] assms osum_wf[of "r - Id" "r' - Id"] - wf_subset[of "(r - Id) +o (r' - Id)" "(r +o r') - Id"] by auto + using osum_minus_Id[of r r'] assms osum_wf[of "r - Id" "r' - Id"] + wf_subset[of "(r - Id) +o (r' - Id)" "(r +o r') - Id"] by auto next have 1: "wf (Inl ` Field r \ Inr ` Field r')" by (rule wf_Int_Times) auto case True @@ -131,13 +132,9 @@ qed lemma osum_Well_order: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "Well_order (r +o r')" -proof- - have "Total r \ Total r'" using WELL WELL' by (auto simp add: order_on_defs) - thus ?thesis using assms unfolding well_order_on_def - using osum_Linear_order osum_wf_Id by blast -qed + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "Well_order (r +o r')" + by (meson WELL WELL' osum_Linear_order osum_wf_Id well_order_on_def wo_rel.TOTAL wo_rel.intro) lemma osum_embedL: assumes WELL: "Well_order r" and WELL': "Well_order r'" @@ -162,10 +159,10 @@ unfolding dir_image_def map_prod_def by auto lemma map_prod_ordIso: "\Well_order r; inj_on f (Field r)\ \ map_prod f f ` r =o r" - unfolding dir_image_alt[symmetric] by (rule ordIso_symmetric[OF dir_image_ordIso]) + by (metis dir_image_alt dir_image_ordIso ordIso_symmetric) definition oprod :: "'a rel \ 'b rel \ ('a \ 'b) rel" (infixr "*o" 80) -where "r *o r' = {((x1, y1), (x2, y2)). + where "r *o r' = {((x1, y1), (x2, y2)). (((y1, y2) \ r' - Id \ x1 \ Field r \ x2 \ Field r) \ ((y1, y2) \ Restr Id (Field r') \ (x1, x2) \ r))}" @@ -178,30 +175,7 @@ lemma oprod_trans: assumes "trans r" "trans r'" "antisym r" "antisym r'" shows "trans (r *o r')" -proof(unfold trans_def, safe) - fix x y z assume *: "(x, y) \ r *o r'" "(y, z) \ r *o r'" - thus "(x, z) \ r *o r'" - unfolding oprod_def - apply safe - apply (metis assms(2) transE) - apply (metis assms(2) transE) - apply (metis assms(2) transE) - apply (metis assms(4) antisymD) - apply (metis assms(4) antisymD) - apply (metis assms(2) transE) - apply (metis assms(4) antisymD) - apply (metis Field_def Range_iff Un_iff) - apply (metis Field_def Range_iff Un_iff) - apply (metis Field_def Range_iff Un_iff) - apply (metis Field_def Domain_iff Un_iff) - apply (metis Field_def Domain_iff Un_iff) - apply (metis Field_def Domain_iff Un_iff) - apply (metis assms(1) transE) - apply (metis assms(1) transE) - apply (metis assms(1) transE) - apply (metis assms(1) transE) - done -qed + using assms by (clarsimp simp: trans_def antisym_def oprod_def) (metis FieldI1 FieldI2) lemma oprod_Preorder: "\Preorder r; Preorder r'; antisym r; antisym r'\ \ Preorder (r *o r')" unfolding preorder_on_def using oprod_Refl oprod_trans by blast @@ -219,9 +193,9 @@ unfolding linear_order_on_def using oprod_Partial_order oprod_Total by blast lemma oprod_wf: -assumes WF: "wf r" and WF': "wf r'" -shows "wf (r *o r')" -unfolding wf_eq_minimal2 unfolding Field_oprod + assumes WF: "wf r" and WF': "wf r'" + shows "wf (r *o r')" + unfolding wf_eq_minimal2 unfolding Field_oprod proof(intro allI impI, elim conjE) fix A assume *: "A \ Field r \ Field r'" and **: "A \ {}" then obtain y where y: "y \ snd ` A" "\y'\snd ` A. (y', y) \ r'" @@ -285,22 +259,17 @@ proof(cases "r \ Id \ r' \ Id") case False thus ?thesis - using oprod_minus_Id[of r r'] assms oprod_wf[of "r - Id" "r' - Id"] - wf_subset[of "(r - Id) *o (r' - Id)" "(r *o r') - Id"] by auto + by (meson TOT TOT' WF WF' oprod_minus_Id oprod_wf wf_subset) next case True thus ?thesis using wf_subset[OF wf_extend_oprod1[OF WF'] oprod_minus_Id1] - wf_subset[OF wf_extend_oprod2[OF WF] oprod_minus_Id2] by auto + wf_subset[OF wf_extend_oprod2[OF WF] oprod_minus_Id2] by auto qed lemma oprod_Well_order: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "Well_order (r *o r')" -proof- - have "Total r \ Total r'" using WELL WELL' by (auto simp add: order_on_defs) - thus ?thesis using assms unfolding well_order_on_def - using oprod_Linear_order oprod_wf_Id by blast -qed + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "Well_order (r *o r')" + by (meson WELL WELL' linear_order_on_def oprod_Linear_order oprod_wf_Id well_order_on_def) lemma oprod_embed: assumes WELL: "Well_order r" and WELL': "Well_order r'" and "r' \ {}" @@ -344,14 +313,8 @@ lemma fun_unequal_in_support: assumes "f \ g" "f \ Func A B" "g \ Func A C" - shows "(support z A f \ support z A g) \ {a. f a \ g a} \ {}" (is "?L \ ?R \ {}") -proof - - from assms(1) obtain x where x: "f x \ g x" by blast - hence "x \ ?R" by simp - moreover from assms(2-3) x have "x \ A" unfolding Func_def by fastforce - with x have "x \ ?L" unfolding support_def by auto - ultimately show ?thesis by auto -qed + shows "(support z A f \ support z A g) \ {a. f a \ g a} \ {}" + using assms by (simp add: Func_def support_def disjoint_iff fun_eq_iff) metis definition fin_support where "fin_support z A = {f. finite (support z A f)}" @@ -381,18 +344,18 @@ begin definition isMaxim :: "'a set \ 'a \ bool" -where "isMaxim A b \ b \ A \ (\a \ A. (a,b) \ r)" + where "isMaxim A b \ b \ A \ (\a \ A. (a,b) \ r)" definition maxim :: "'a set \ 'a" -where "maxim A \ THE b. isMaxim A b" + where "maxim A \ THE b. isMaxim A b" lemma isMaxim_unique[intro]: "\isMaxim A x; isMaxim A y\ \ x = y" unfolding isMaxim_def using antisymD[OF ANTISYM, of x y] by auto lemma maxim_isMaxim: "\finite A; A \ {}; A \ Field r\ \ isMaxim A (maxim A)" -unfolding maxim_def + unfolding maxim_def proof (rule theI', rule ex_ex1I[OF _ isMaxim_unique, rotated], assumption+, - induct A rule: finite_induct) + induct A rule: finite_induct) case (insert x A) thus ?case proof (cases "A = {}") @@ -426,9 +389,10 @@ show ?thesis proof (cases "(x, maxim A) \ r") case True - with *(2) have "isMaxim (insert x A) (maxim A)" unfolding isMaxim_def - using transD[OF TRANS, of _ x "maxim A"] by blast - with *(1) True show ?thesis unfolding max2_def by (metis isMaxim_unique) + with *(2) have "isMaxim (insert x A) (maxim A)" + by (simp add: isMaxim_def) + with *(1) True show ?thesis + unfolding max2_def by (metis isMaxim_unique) next case False hence "(maxim A, x) \ r" by (metis *(2) assms(3,4) in_mono in_notinI isMaxim_def) @@ -453,8 +417,8 @@ next case False hence "(maxim B, maxim A) \ r" by (metis *(2,3) assms(3,6) in_mono in_notinI isMaxim_def) - with *(2,3) have "isMaxim (A \ B) (maxim A)" unfolding isMaxim_def - using transD[OF TRANS, of _ "maxim B" "maxim A"] by blast + with *(2,3) have "isMaxim (A \ B) (maxim A)" + by (metis "*"(1) False Un_iff isMaxim_def isMaxim_unique) with *(1) False show ?thesis unfolding max2_def by (metis isMaxim_unique) qed qed @@ -462,7 +426,7 @@ lemma maxim_insert_zero: assumes "finite A" "A \ {}" "A \ Field r" shows "maxim (insert zero A) = maxim A" -using assms zero_in_Field maxim_in[OF assms] by (subst maxim_insert[unfolded max2_def]) auto + using assms finite.cases in_mono max2_def maxim_in maxim_insert subset_empty zero_in_Field zero_smallest by fastforce lemma maxim_equality: "isMaxim A x \ maxim A = x" unfolding maxim_def by (rule the_equality) auto @@ -495,7 +459,7 @@ locale wo_rel2 = fixes r s assumes rWELL: "Well_order r" - and sWELL: "Well_order s" + and sWELL: "Well_order s" begin interpretation r: wo_rel r by unfold_locales (rule rWELL) @@ -510,7 +474,7 @@ lemma max_fun_diff_alt: "s.max_fun_diff f g = s.maxim ((SUPP f \ SUPP g) \ {a. f a \ g a})" - unfolding s.max_fun_diff_def fun_diff_alt .. + unfolding s.max_fun_diff_def fun_diff_alt .. lemma isMaxim_max_fun_diff: "\f \ g; f \ FINFUNC; g \ FINFUNC\ \ s.isMaxim {a \ Field s. f a \ g a} (s.max_fun_diff f g)" @@ -552,13 +516,7 @@ hence "(x, ?fg) \ s" proof (cases "x = ?fg") case False show ?thesis - proof (rule ccontr) - assume "(x, ?fg) \ s" - with max_fun_diff_in[OF fg f g] x False have *: "(?fg, x) \ s" by (blast intro: s.in_notinI) - hence "f x = g x" by (rule max_fun_diff_le_eq[OF _ fg f g False]) - moreover have "g x = h x" using max_fun_diff_le_eq[OF _ gh g h] False True * by simp - ultimately show False using x by simp - qed + by (metis (mono_tags, lifting) True assms(5-7) max_fun_diff_max mem_Collect_eq x) qed (simp add: refl_onD[OF s.REFL]) } ultimately have "s.isMaxim {a \ Field s. f a \ h a} ?fg" @@ -572,7 +530,7 @@ case True hence *: "f ?gh = g ?gh" by (rule max_fun_diff_le_eq[OF _ fg f g *[symmetric]]) hence "s.isMaxim {a \ Field s. f a \ h a} ?gh" using isMaxim_max_fun_diff[OF gh g h] - isMaxim_max_fun_diff[OF fg f g] transD[OF s.TRANS _ True] + isMaxim_max_fun_diff[OF fg f g] transD[OF s.TRANS _ True] unfolding s.isMaxim_def by auto hence "?fh = ?gh" using isMaxim_max_fun_diff[OF fh f h] by blast thus ?thesis using True unfolding s.max2_def by simp @@ -582,7 +540,7 @@ by (blast intro: s.in_notinI) hence *: "g ?fg = h ?fg" by (rule max_fun_diff_le_eq[OF _ gh g h *]) hence "s.isMaxim {a \ Field s. f a \ h a} ?fg" using isMaxim_max_fun_diff[OF gh g h] - isMaxim_max_fun_diff[OF fg f g] True transD[OF s.TRANS, of _ _ ?fg] + isMaxim_max_fun_diff[OF fg f g] True transD[OF s.TRANS, of _ _ ?fg] unfolding s.isMaxim_def by auto hence "?fh = ?fg" using isMaxim_max_fun_diff[OF fh f h] by blast thus ?thesis using False unfolding s.max2_def by simp @@ -603,8 +561,8 @@ proof (unfold trans_def, safe) fix f g h :: "'b \ 'a" let ?fg = "s.max_fun_diff f g" - and ?gh = "s.max_fun_diff g h" - and ?fh = "s.max_fun_diff f h" + and ?gh = "s.max_fun_diff g h" + and ?fh = "s.max_fun_diff f h" assume oexp: "(f, g) \ oexp" "(g, h) \ oexp" thus "(f, h) \ oexp" proof (cases "f = g \ g = h") @@ -619,8 +577,8 @@ proof (cases "?fg = ?gh \ f ?fg \ h ?gh") case True show ?thesis using max_fun_diff_max2[of f g h, OF True] * \f \ h\ max_fun_diff_in - r.max2_iff[OF FINFUNCD FINFUNCD] r.max2_equals1[OF FINFUNCD FINFUNCD] max_fun_diff_le_eq - s.in_notinI[OF disjI1] unfolding oexp_def Let_def s.max2_def mem_Collect_eq by safe metis + r.max2_iff[OF FINFUNCD FINFUNCD] r.max2_equals1[OF FINFUNCD FINFUNCD] max_fun_diff_le_eq + s.in_notinI[OF disjI1] unfolding oexp_def Let_def s.max2_def mem_Collect_eq by safe metis next case False with * show ?thesis unfolding oexp_def Let_def using antisymD[OF r.ANTISYM, of "g ?gh" "h ?gh"] max_fun_diff_in[of g h] by auto @@ -667,13 +625,7 @@ lemma const_least: assumes "Field r \ {}" "f \ FINFUNC" shows "(const, f) \ oexp" -proof (cases "f = const") - case True thus ?thesis using refl_onD[OF oexp_Refl] assms(2) unfolding Field_oexp by auto -next - case False - with assms show ?thesis using max_fun_diff_in[of f const] - unfolding oexp_def Let_def by (auto intro: r.zero_smallest FinFuncD simp: s.max_fun_diff_commute) -qed + using assms const_FINFUNC max_fun_diff max_fun_diff_in oexp_def by fastforce lemma support_not_const: assumes "F \ FINFUNC" and "const \ F" @@ -693,15 +645,15 @@ qed lemma maxim_isMaxim_support: - assumes f: "F \ FINFUNC" and "const \ F" + assumes "F \ FINFUNC" and "const \ F" shows "\f \ F. s.isMaxim (SUPP f) (s.maxim (SUPP f))" - using support_not_const[OF assms] by (auto intro!: s.maxim_isMaxim) + using assms s.maxim_isMaxim support_not_const by force lemma oexp_empty2: "Field s = {} \ oexp = {(\x. undefined, \x. undefined)}" unfolding oexp_def FinFunc_def fin_support_def support_def by auto lemma oexp_empty: "\Field r = {}; Field s \ {}\ \ oexp = {}" - unfolding oexp_def FinFunc_def Let_def by auto + using FINFUNCD oexp_def by auto lemma fun_upd_FINFUNC: "\f \ FINFUNC; x \ Field s; y \ Field r\ \ f(x := y) \ FINFUNC" unfolding FinFunc_def Func_def fin_support_def @@ -734,7 +686,7 @@ have const[simp]: "\F. \const \ F; F \ FINFUNC\ \ \f0\F. \f\F. (f0, f) \ oexp" using const_least[OF Fields(2)] by auto show ?thesis - unfolding Linear_order_wf_diff_Id[OF oexp_Linear_order] Field_oexp + unfolding Linear_order_wf_diff_Id[OF oexp_Linear_order] Field_oexp proof (intro allI impI) fix A assume A: "A \ FINFUNC" "A \ {}" { fix y F @@ -758,7 +710,7 @@ hence zField: "z \ Field s" unfolding Field_def by auto define x0 where "x0 = r.minim {f z | f. f \ F \ z = s.maxim (SUPP f)}" from F(1,2) maxF(1) SUPPF zmin - have x0min: "r.isMinim {f z | f. f \ F \ z = s.maxim (SUPP f)} x0" + have x0min: "r.isMinim {f z | f. f \ F \ z = s.maxim (SUPP f)} x0" unfolding x0_def s.isMaxim_def s.isMinim_def by (blast intro!: r.minim_isMinim FinFuncD[of _ r s]) with maxF(1) SUPPF F(1) have x0Field: "x0 \ Field r" @@ -770,13 +722,8 @@ from zmin x0min have "G \ {}" unfolding G_def z_def s.isMinim_def r.isMinim_def by blast have GF: "G \ (\f. f(z := r.zero)) ` F" unfolding G_def by auto have "G \ fin_support r.zero (Field s)" - unfolding FinFunc_def fin_support_def proof safe - fix g assume "g \ G" - with GF obtain f where f: "f \ F" "g = f(z := r.zero)" by auto - with SUPPF have "finite (SUPP f)" by blast - with f show "finite (SUPP g)" - by (elim finite_subset[rotated]) (auto simp: support_def) - qed + unfolding FinFunc_def fin_support_def + using F(1) FinFunc_def G_def fin_support_def by fastforce moreover from F GF zField have "G \ Func (Field s) (Field r)" using Func_upd[of _ "Field s" "Field r" z r.zero] unfolding FinFunc_def by auto ultimately have G: "G \ FINFUNC" unfolding FinFunc_def by blast @@ -799,7 +746,7 @@ unfolding z by (intro s.maxim_mono) auto } moreover from y'min have "\g. g \ G \ (y', s.maxim (SUPP g)) \ s" - unfolding s.isMinim_def by auto + unfolding s.isMinim_def by auto ultimately have "y' \ z" "(y', z) \ s" using maxG unfolding s.isMinim_def s.isMaxim_def by auto with zy have "y' \ y" "(y', y) \ s" using antisymD[OF s.ANTISYM] transD[OF s.TRANS] @@ -814,7 +761,7 @@ with x0notzero zField have SUPP: "SUPP f0 = SUPP g0 \ {z}" unfolding support_def by auto from g0z have f0z: "f0(z := r.zero) = g0" unfolding f0_def fun_upd_upd by auto have f0: "f0 \ F" using x0min g0(1) - Func_elim[OF subsetD[OF subset_trans[OF F(1)[unfolded FinFunc_def] Int_lower1]] zField] + Func_elim[OF subsetD[OF subset_trans[OF F(1)[unfolded FinFunc_def] Int_lower1]] zField] unfolding f0_def r.isMinim_def G_def by (force simp: fun_upd_idem) from g0(1) maxF(1) have maxf0: "s.maxim (SUPP f0) = z" unfolding SUPP G_def by (intro s.maxim_equality) (auto simp: s.isMaxim_def) @@ -834,7 +781,7 @@ hence oexp: "(f0(z := r.zero, z := x0), f(z := r.zero, z := x0)) \ oexp" by (elim fun_upd_same_oexp[OF _ _ zField x0Field]) simp with f F(1) x0min True - have "(f(z := x0), f) \ oexp" unfolding G_def r.isMinim_def + have "(f(z := x0), f) \ oexp" unfolding G_def r.isMinim_def by (intro fun_upd_smaller_oexp[OF _ zField x0Field]) auto with oexp show ?thesis using transD[OF oexp_trans, of f0 "f(z := x0)" f] unfolding f0_def by auto @@ -868,12 +815,12 @@ have "f (s.maxim (SUPP f)) \ r.zero" using bspec[OF maxF(1) f, unfolded s.isMaxim_def] by (auto simp: support_def) with f0f * f f0 maxf0 SUPPF - have "s.max_fun_diff f0 f = s.maxim (SUPP f0 \ SUPP f)" + have "s.max_fun_diff f0 f = s.maxim (SUPP f0 \ SUPP f)" unfolding max_fun_diff_alt using s.maxim_Un[of "SUPP f0" "SUPP f"] by (intro s.maxim_Int) (auto simp: s.max2_def) moreover have "s.maxim (SUPP f0 \ SUPP f) = s.maxim (SUPP f)" - using s.maxim_Un[of "SUPP f0" "SUPP f"] * maxf0 SUPPF f0 f - by (auto simp: s.max2_def) + using s.maxim_Un[of "SUPP f0" "SUPP f"] * maxf0 SUPPF f0 f + by (auto simp: s.max2_def) ultimately show ?thesis using f f0 F(1) maxF(1) SUPPF unfolding oexp_def Let_def by (fastforce simp: s.isMaxim_def intro!: r.zero_smallest FINFUNCD) qed @@ -883,13 +830,9 @@ qed simp qed qed - } note * = mp[OF this] - from A(2) obtain f where f: "f \ A" by blast - with A(1) show "\a\A. \a'\A. (a, a') \ oexp" - proof (cases "f = const") - case False with f A(1) show ?thesis using maxim_isMaxim_support[of "{f}"] - by (intro *[of _ "s.maxim (SUPP f)"]) (auto simp: s.isMaxim_def support_def) - qed simp + } + with A show "\a\A. \a'\A. (a, a') \ oexp" + by blast qed qed @@ -899,8 +842,7 @@ interpretation o: wo_rel oexp by unfold_locales (rule oexp_Well_order) lemma zero_oexp: "Field r \ {} \ o.zero = const" - by (rule sym[OF o.leq_zero_imp[OF const_least]]) - (auto intro!: o.zero_in_Field[unfolded Field_oexp] dest!: const_FINFUNC) + by (metis Field_oexp const_FINFUNC const_least o.Field_ofilter o.equals_minim o.ofilter_def o.zero_def) end @@ -925,8 +867,8 @@ by (auto dest: well_order_on_domain) lemma ozero_ordLeq: -assumes "Well_order r" shows "ozero \o r" -using assms unfolding ozero_def ordLeq_def embed_def[abs_def] under_def by auto + assumes "Well_order r" shows "ozero \o r" + using assms unfolding ozero_def ordLeq_def embed_def[abs_def] under_def by auto definition "oone = {((),())}" @@ -985,8 +927,8 @@ lemma osum_cong: assumes "t =o u" and "r =o s" shows "t +o r =o u +o s" -using ordIso_transitive[OF osum_congL[OF assms(1)] osum_congR[OF assms(2)]] - assms[unfolded ordIso_def] by auto + using ordIso_transitive[OF osum_congL[OF assms(1)] osum_congR[OF assms(2)]] + assms[unfolded ordIso_def] by auto lemma Well_order_empty[simp]: "Well_order {}" unfolding Field_empty by (rule well_order_on_empty) @@ -1005,8 +947,8 @@ shows "{} ^o r = {}" proof - from assms(2) have "Field r \ {}" unfolding Field_def by auto - thus ?thesis unfolding oexp_def[OF Well_order_empty assms(1)] FinFunc_def fin_support_def support_def - by auto + thus ?thesis + by (simp add: assms(1) wo_rel2.intro wo_rel2.oexp_empty) qed lemma oprod_zero[simp]: "{} *o r = {}" "r *o {} = {}" @@ -1051,8 +993,8 @@ lemma oprod_cong: assumes "t =o u" and "r =o s" shows "t *o r =o u *o s" -using ordIso_transitive[OF oprod_congL[OF assms(1)] oprod_congR[OF assms(2)]] - assms[unfolded ordIso_def] by auto + using ordIso_transitive[OF oprod_congL[OF assms(1)] oprod_congR[OF assms(2)]] + assms[unfolded ordIso_def] by auto lemma Field_singleton[simp]: "Field {(z,z)} = {z}" by (metis well_order_on_Field well_order_on_singleton) @@ -1155,7 +1097,7 @@ let ?f = "map_sum f id" from f have "\a\Field (r +o t). ?f a \ Field (s +o t) \ ?f ` underS (r +o t) a \ underS (s +o t) (?f a)" - unfolding Field_osum underS_def by (fastforce simp: osum_def) + unfolding Field_osum underS_def by (fastforce simp: osum_def) thus ?thesis unfolding ordLeq_def2 by (auto intro: osum_Well_order r s t) qed @@ -1224,7 +1166,7 @@ let ?f = "map_prod f id" from f have "\a\Field (r *o t). ?f a \ Field (s *o t) \ ?f ` underS (r *o t) a \ underS (s *o t) (?f a)" - unfolding Field_oprod underS_def unfolding map_prod_def oprod_def by auto + unfolding Field_oprod underS_def unfolding map_prod_def oprod_def by auto thus ?thesis unfolding ordLeq_def2 by (auto intro: oprod_Well_order r s t) qed @@ -1266,8 +1208,7 @@ qed lemma ozero_oexp: "\ (s =o ozero) \ ozero ^o s =o ozero" - unfolding oexp_def[OF ozero_Well_order s] FinFunc_def - by simp (metis Func_emp2 bot.extremum_uniqueI emptyE well_order_on_domain s subrelI) + by (fastforce simp add: oexp_def[OF ozero_Well_order s] FinFunc_def Func_def intro: FieldI1) lemma oone_oexp: "oone ^o s =o oone" (is "?L =o ?R") by (rule oone_ordIso_oexp[OF ordIso_reflexive[OF oone_Well_order] s]) @@ -1315,17 +1256,17 @@ "let m = s.max_fun_diff g h in (g m, h m) \ r" hence "g \ h" by auto note max_fun_diff_in = rs.max_fun_diff_in[OF \g \ h\ gh(1,2)] - and max_fun_diff_max = rs.max_fun_diff_max[OF \g \ h\ gh(1,2)] + and max_fun_diff_max = rs.max_fun_diff_max[OF \g \ h\ gh(1,2)] with *(4) invff *(2) have "t.max_fun_diff (F g) (F h) = f (s.max_fun_diff g h)" unfolding t.max_fun_diff_def compat_def by (intro t.maxim_equality) (auto simp: t.isMaxim_def F_def dest: injfD) with gh invff max_fun_diff_in - show "let m = t.max_fun_diff (F g) (F h) in (F g m, F h m) \ r" + show "let m = t.max_fun_diff (F g) (F h) in (F g m, F h m) \ r" unfolding F_def Let_def by (auto simp: dest: injfD) qed moreover from FLR have "ofilter ?R (F ` Field ?L)" - unfolding rexpt.ofilter_def under_def rs.Field_oexp rt.Field_oexp unfolding rt.oexp_def + unfolding rexpt.ofilter_def under_def rs.Field_oexp rt.Field_oexp unfolding rt.oexp_def proof (safe elim!: imageI) fix g h assume gh: "g \ FinFunc r s" "h \ FinFunc r t" "F g \ FinFunc r t" "let m = t.max_fun_diff h (F g) in (h m, F g m) \ r" @@ -1345,13 +1286,13 @@ proof (rule ccontr) assume "(t.max_fun_diff h (F g), z) \ t" hence "(z, t.max_fun_diff h (F g)) \ t" using t.in_notinI[of "t.max_fun_diff h (F g)" z] - z max_Field by auto + z max_Field by auto hence "z \ f ` Field s" using *(3) max_f_Field unfolding t.ofilter_def under_def by fastforce with z show False by blast qed hence "h z = r.zero" using rt.max_fun_diff_le_eq[OF _ False gh(2,3), of z] - z max_f_Field unfolding F_def by auto + z max_f_Field unfolding F_def by auto } note ** = this with *(3) gh(2) have "h = F (\x. if x \ Field s then h (f x) else undefined)" using invff unfolding F_def fun_eq_iff FinFunc_def Func_def Let_def t.ofilter_def under_def by auto @@ -1437,8 +1378,8 @@ by (subst t.max_fun_diff_def, intro t.maxim_equality) (auto simp: t.isMaxim_def intro: inj_onD[OF inj] intro!: rt.max_fun_diff_max) with Field_fg Field_fh hg fz f_underS compat neq have "(?f h, ?f g) \ st.oexp" - using rt.max_fun_diff[OF \h \ g\] rt.max_fun_diff_in[OF \h \ g\] unfolding st.Field_oexp - unfolding rt.oexp_def st.oexp_def Let_def compat_def by auto + using rt.max_fun_diff[OF \h \ g\] rt.max_fun_diff_in[OF \h \ g\] unfolding st.Field_oexp + unfolding rt.oexp_def st.oexp_def Let_def compat_def by auto with neq show "?f h \ underS (s ^o t) (?f g)" unfolding underS_def by auto qed ultimately have "?f g \ Field (s ^o t) \ ?f ` underS (r ^o t) g \ underS (s ^o t) (?f g)" @@ -1460,7 +1401,7 @@ x: "x \ Field r" "r.zero \ Field r" "x \ r.zero" unfolding ordLess_def oone_def embedS_def[abs_def] bij_betw_def embed_def under_def by (auto simp: image_def) - (metis (lifting) equals0D mem_Collect_eq r.zero_in_Field singletonI) + (metis (lifting) equals0D mem_Collect_eq r.zero_in_Field singletonI) let ?f = "\a b. if b \ Field s then if b = a then x else r.zero else undefined" from x(3) have SUPP: "\y. y \ Field s \ rs.SUPP (?f y) = {y}" unfolding support_def by auto { fix y assume y: "y \ Field s" @@ -1553,7 +1494,7 @@ interpret rt: wo_rel2 r t by unfold_locales (rule r, rule t) let ?f = "\(f, g). case_sum f g" have "bij_betw ?f (Field ?L) (Field ?R)" - unfolding bij_betw_def rst.Field_oexp rs.Field_oexp rt.Field_oexp Field_oprod proof (intro conjI) + unfolding bij_betw_def rst.Field_oexp rs.Field_oexp rt.Field_oexp Field_oprod proof (intro conjI) show "inj_on ?f (FinFunc r s \ FinFunc r t)" unfolding inj_on_def by (auto simp: fun_eq_iff split: sum.splits) show "?f ` (FinFunc r s \ FinFunc r t) = FinFunc r (s +o t)" @@ -1567,7 +1508,7 @@ moreover have "compat ?L ?R ?f" unfolding compat_def rst.Field_oexp rs.Field_oexp rt.Field_oexp oprod_def unfolding rst.oexp_def Let_def rs.oexp_def rt.oexp_def - by (fastforce simp: Field_osum FinFunc_osum o_def split: sum.splits + by (fastforce simp: Field_osum FinFunc_osum o_def split: sum.splits dest: max_fun_diff_eq_Inl max_fun_diff_eq_Inr) ultimately have "iso ?L ?R ?f" using r s t by (subst iso_iff3) (auto intro: oexp_Well_order oprod_Well_order osum_Well_order) @@ -1638,8 +1579,8 @@ interpret rs: wo_rel2 r s by unfold_locales (rule r, rule s) interpret rst: wo_rel2 "r ^o s" t by unfold_locales (rule oexp_Well_order[OF r s], rule t) from fun_unequal_in_support[OF assms(2), of "Field (s *o t)" "Field r" "Field r"] assms(3,4) - have diff1: "rev_curr f \ rev_curr g" - "rev_curr f \ FinFunc (r ^o s) t" "rev_curr g \ FinFunc (r ^o s) t" using rev_curr_FinFunc[OF Field] + have diff1: "rev_curr f \ rev_curr g" + "rev_curr f \ FinFunc (r ^o s) t" "rev_curr g \ FinFunc (r ^o s) t" using rev_curr_FinFunc[OF Field] unfolding fun_eq_iff rev_curr_def[abs_def] FinFunc_def support_def Field_oprod by auto fast hence diff2: "rev_curr f m \ rev_curr g m" "rev_curr f m \ FinFunc r s" "rev_curr g m \ FinFunc r s" @@ -1653,7 +1594,7 @@ next assume "f (s.max_fun_diff (rev_curr f m) (rev_curr g m), m) = g (s.max_fun_diff (rev_curr f m) (rev_curr g m), m)" - (is "f (?x, m) = g (?x, m)") + (is "f (?x, m) = g (?x, m)") hence "rev_curr f m ?x = rev_curr g m ?x" unfolding rev_curr_def by auto with rs.max_fun_diff[OF diff2] show False by auto next @@ -1674,14 +1615,14 @@ proof (cases "s = {} \ t = {}") case True with \r = {}\ show ?thesis by (auto simp: oexp_empty[OF oexp_Well_order[OF Well_order_empty s]] - intro!: ordIso_transitive[OF ordIso_symmetric[OF oone_ordIso] oone_ordIso] + intro!: ordIso_transitive[OF ordIso_symmetric[OF oone_ordIso] oone_ordIso] ordIso_transitive[OF oone_ordIso_oexp[OF ordIso_symmetric[OF oone_ordIso] t] oone_ordIso]) next - case False - hence "s *o t \ {}" unfolding oprod_def Field_def by fastforce - with False show ?thesis - using \r = {}\ ozero_ordIso - by (auto simp add: s t oprod_Well_order ozero_def) + case False + hence "s *o t \ {}" unfolding oprod_def Field_def by fastforce + with False show ?thesis + using \r = {}\ ozero_ordIso + by (auto simp add: s t oprod_Well_order ozero_def) qed next case False @@ -1692,7 +1633,7 @@ interpret rs: wo_rel2 r s by unfold_locales (rule r, rule s) interpret rst: wo_rel2 "r ^o s" t by unfold_locales (rule oexp_Well_order[OF r s], rule t) have bij: "bij_betw rev_curr (Field ?L) (Field ?R)" - unfolding bij_betw_def r_st.Field_oexp rst.Field_oexp Field_oprod proof (intro conjI) + unfolding bij_betw_def r_st.Field_oexp rst.Field_oexp Field_oprod proof (intro conjI) show "inj_on rev_curr (FinFunc r (s *o t))" unfolding inj_on_def FinFunc_def Func_def Field_oprod rs.Field_oexp rev_curr_def[abs_def] by (auto simp: fun_eq_iff) metis @@ -1700,16 +1641,16 @@ qed moreover have "compat ?L ?R rev_curr" - unfolding compat_def proof safe + unfolding compat_def proof safe fix fg1 fg2 assume fg: "(fg1, fg2) \ r ^o (s *o t)" show "(rev_curr fg1, rev_curr fg2) \ r ^o s ^o t" proof (cases "fg1 = fg2") assume "fg1 \ fg2" with fg show ?thesis - using rst.max_fun_diff_in[of "rev_curr fg1" "rev_curr fg2"] - max_fun_diff_oprod[OF Field, of fg1 fg2] rev_curr_FinFunc[OF Field, symmetric] - unfolding r_st.Field_oexp rs.Field_oexp rst.Field_oexp unfolding r_st.oexp_def rst.oexp_def - by (auto simp: rs.oexp_def Let_def) (auto simp: rev_curr_def[abs_def]) + using rst.max_fun_diff_in[of "rev_curr fg1" "rev_curr fg2"] + max_fun_diff_oprod[OF Field, of fg1 fg2] rev_curr_FinFunc[OF Field, symmetric] + unfolding r_st.Field_oexp rs.Field_oexp rst.Field_oexp unfolding r_st.oexp_def rst.oexp_def + by (auto simp: rs.oexp_def Let_def) (auto simp: rev_curr_def[abs_def]) next assume "fg1 = fg2" with fg bij show ?thesis unfolding r_st.Field_oexp rs.Field_oexp rst.Field_oexp bij_betw_def diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Cardinals/Wellorder_Constructions.thy --- a/src/HOL/Cardinals/Wellorder_Constructions.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Cardinals/Wellorder_Constructions.thy Sun Jan 15 20:00:44 2023 +0100 @@ -8,8 +8,8 @@ section \Constructions on Wellorders\ theory Wellorder_Constructions -imports - Wellorder_Embedding Order_Union + imports + Wellorder_Embedding Order_Union begin unbundle cardinal_syntax @@ -21,254 +21,158 @@ Func_empty[simp] Func_is_emp[simp] -lemma Func_emp2[simp]: "A \ {} \ Func A {} = {}" by auto - - -subsection \Restriction to a set\ - -lemma Restr_incr2: -"r <= r' \ Restr r A <= Restr r' A" -by blast - -lemma Restr_incr: -"\r \ r'; A \ A'\ \ Restr r A \ Restr r' A'" -by blast - -lemma Restr_Int: -"Restr (Restr r A) B = Restr r (A Int B)" -by blast - -lemma Restr_iff: "(a,b) \ Restr r A = (a \ A \ b \ A \ (a,b) \ r)" -by (auto simp add: Field_def) - -lemma Restr_subset1: "Restr r A \ r" -by auto - -lemma Restr_subset2: "Restr r A \ A \ A" -by auto - -lemma wf_Restr: -"wf r \ wf(Restr r A)" -using Restr_subset by (elim wf_subset) simp - -lemma Restr_incr1: -"A \ B \ Restr r A \ Restr r B" -by blast subsection \Order filters versus restrictions and embeddings\ -lemma ofilter_Restr: -assumes WELL: "Well_order r" and - OFA: "ofilter r A" and OFB: "ofilter r B" and SUB: "A \ B" -shows "ofilter (Restr r B) A" -proof- - let ?rB = "Restr r B" - have Well: "wo_rel r" unfolding wo_rel_def using WELL . - hence Refl: "Refl r" by (auto 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 (auto simp add: Well_order_Restr wo_rel_def) - (* Main proof *) - show ?thesis - proof(auto simp add: WellB wo_rel.ofilter_def) - fix a assume "a \ A" - hence "a \ Field r \ a \ B" using assms Well - by (auto simp add: wo_rel.ofilter_def) - with Field show "a \ Field(Restr r B)" by auto - next - fix a b assume *: "a \ A" and "b \ under (Restr r B) a" - hence "b \ under r a" - using WELL OFB SUB ofilter_Restr_under[of r B a] by auto - thus "b \ A" using * Well OFA by(auto simp add: wo_rel.ofilter_def) - qed -qed - lemma ofilter_subset_iso: -assumes WELL: "Well_order r" and - OFA: "ofilter r A" and OFB: "ofilter r B" -shows "(A = B) = iso (Restr r A) (Restr r B) id" -using assms -by (auto simp add: ofilter_subset_embedS_iso) + assumes WELL: "Well_order r" and + OFA: "ofilter r A" and OFB: "ofilter r B" + shows "(A = B) = iso (Restr r A) (Restr r B) id" + using assms by (auto simp add: ofilter_subset_embedS_iso) subsection \Ordering the well-orders by existence of embeddings\ corollary ordLeq_refl_on: "refl_on {r. Well_order r} ordLeq" -using ordLeq_reflexive unfolding ordLeq_def refl_on_def -by blast + using ordLeq_reflexive unfolding ordLeq_def refl_on_def + by blast corollary ordLeq_trans: "trans ordLeq" -using trans_def[of ordLeq] ordLeq_transitive by blast + using trans_def[of ordLeq] ordLeq_transitive by blast corollary ordLeq_preorder_on: "preorder_on {r. Well_order r} ordLeq" -by(auto simp add: preorder_on_def ordLeq_refl_on ordLeq_trans) + by(auto simp add: preorder_on_def ordLeq_refl_on ordLeq_trans) corollary ordIso_refl_on: "refl_on {r. Well_order r} ordIso" -using ordIso_reflexive unfolding refl_on_def ordIso_def -by blast + using ordIso_reflexive unfolding refl_on_def ordIso_def + by blast corollary ordIso_trans: "trans ordIso" -using trans_def[of ordIso] ordIso_transitive by blast + using trans_def[of ordIso] ordIso_transitive by blast corollary ordIso_sym: "sym ordIso" -by (auto simp add: sym_def ordIso_symmetric) + by (auto simp add: sym_def ordIso_symmetric) corollary ordIso_equiv: "equiv {r. Well_order r} ordIso" -by (auto simp add: equiv_def ordIso_sym ordIso_refl_on ordIso_trans) + by (auto simp add: equiv_def ordIso_sym ordIso_refl_on ordIso_trans) lemma ordLess_Well_order_simp[simp]: -assumes "r Well_order r'" -using assms unfolding ordLess_def by simp + assumes "r Well_order r'" + using assms unfolding ordLess_def by simp lemma ordIso_Well_order_simp[simp]: -assumes "r =o r'" -shows "Well_order r \ Well_order r'" -using assms unfolding ordIso_def by simp + assumes "r =o r'" + shows "Well_order r \ Well_order r'" + using assms unfolding ordIso_def by simp lemma ordLess_irrefl: "irrefl ordLess" -by(unfold irrefl_def, auto simp add: ordLess_irreflexive) + by(unfold irrefl_def, auto simp add: ordLess_irreflexive) lemma ordLess_or_ordIso: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "r r' r =o r'" -unfolding ordLess_def ordIso_def -using assms embedS_or_iso[of r r'] by auto + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "r r' r =o r'" + unfolding ordLess_def ordIso_def + using assms embedS_or_iso[of r r'] by auto corollary ordLeq_ordLess_Un_ordIso: -"ordLeq = ordLess \ ordIso" -by (auto simp add: ordLeq_iff_ordLess_or_ordIso) - -lemma not_ordLeq_ordLess: -"r \o r' \ \ r' ordIso" + by (auto simp add: ordLeq_iff_ordLess_or_ordIso) lemma ordIso_or_ordLess: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "r =o r' \ r r' r r' o r" -proof- - have "A \ Field r" using assms by (auto simp add: wo_rel_def wo_rel.ofilter_def) - thus ?thesis using assms - by (simp add: ofilter_subset_ordLeq wo_rel.Field_ofilter - wo_rel_def Restr_Field) -qed + assumes "Well_order r" and "ofilter r A" + shows "Restr r A \o r" +by (metis assms inf.orderE ofilter_embed ofilter_subset_ordLeq refl_on_def wo_rel.Field_ofilter wo_rel.REFL wo_rel.intro) corollary under_Restr_ordLeq: -"Well_order r \ Restr r (under r a) \o r" -by (auto simp add: ofilter_ordLeq wo_rel.under_ofilter wo_rel_def) + "Well_order r \ Restr r (under r a) \o r" + by (auto simp add: ofilter_ordLeq wo_rel.under_ofilter wo_rel_def) subsection \Copy via direct images\ lemma Id_dir_image: "dir_image Id f \ Id" -unfolding dir_image_def by auto + unfolding dir_image_def by auto lemma Un_dir_image: -"dir_image (r1 \ r2) f = (dir_image r1 f) \ (dir_image r2 f)" -unfolding dir_image_def by auto + "dir_image (r1 \ r2) f = (dir_image r1 f) \ (dir_image r2 f)" + unfolding dir_image_def by auto lemma Int_dir_image: -assumes "inj_on f (Field r1 \ Field r2)" -shows "dir_image (r1 Int r2) f = (dir_image r1 f) Int (dir_image r2 f)" + assumes "inj_on f (Field r1 \ Field r2)" + shows "dir_image (r1 Int r2) f = (dir_image r1 f) Int (dir_image r2 f)" proof show "dir_image (r1 Int r2) f \ (dir_image r1 f) Int (dir_image r2 f)" - using assms unfolding dir_image_def inj_on_def by auto + using assms unfolding dir_image_def inj_on_def by auto next show "(dir_image r1 f) Int (dir_image r2 f) \ dir_image (r1 Int r2) f" - proof(clarify) - fix a' b' - assume "(a',b') \ dir_image r1 f" "(a',b') \ dir_image r2 f" - then obtain a1 b1 a2 b2 - where 1: "a' = f a1 \ b' = f b1 \ a' = f a2 \ b' = f b2" and - 2: "(a1,b1) \ r1 \ (a2,b2) \ r2" and - 3: "{a1,b1} \ Field r1 \ {a2,b2} \ Field r2" - unfolding dir_image_def Field_def by blast - hence "a1 = a2 \ b1 = b2" using assms unfolding inj_on_def by auto - hence "a' = f a1 \ b' = f b1 \ (a1,b1) \ r1 Int r2 \ (a2,b2) \ r1 Int r2" - using 1 2 by auto - thus "(a',b') \ dir_image (r1 \ r2) f" - unfolding dir_image_def by blast - qed + by (clarsimp simp: dir_image_def) (metis FieldI1 FieldI2 UnCI assms inj_on_def) qed (* More facts on ordinal sum: *) lemma Osum_embed: -assumes FLD: "Field r Int Field r' = {}" and - WELL: "Well_order r" and WELL': "Well_order r'" -shows "embed r (r Osum r') id" + assumes FLD: "Field r Int Field r' = {}" and + WELL: "Well_order r" and WELL': "Well_order r'" + shows "embed r (r Osum r') id" proof- have 1: "Well_order (r Osum r')" - using assms by (auto simp add: Osum_Well_order) + using assms by (auto simp add: Osum_Well_order) moreover have "compat r (r Osum r') id" - unfolding compat_def Osum_def by auto + unfolding compat_def Osum_def by auto moreover have "inj_on id (Field r)" by simp moreover have "ofilter (r Osum r') (Field r)" - using 1 proof(auto simp add: wo_rel_def wo_rel.ofilter_def - Field_Osum under_def) - fix a b assume 2: "a \ Field r" and 3: "(b,a) \ r Osum r'" - moreover - {assume "(b,a) \ r'" - hence "a \ Field r'" using Field_def[of r'] by blast - hence False using 2 FLD by blast - } - moreover - {assume "a \ Field r'" - hence False using 2 FLD by blast - } - ultimately - show "b \ Field r" by (auto simp add: Osum_def Field_def) - qed + using 1 FLD + by (auto simp add: wo_rel_def wo_rel.ofilter_def Osum_def under_def Field_iff disjoint_iff) ultimately show ?thesis - using assms by (auto simp add: embed_iff_compat_inj_on_ofilter) + using assms by (auto simp add: embed_iff_compat_inj_on_ofilter) qed corollary Osum_ordLeq: -assumes FLD: "Field r Int Field r' = {}" and - WELL: "Well_order r" and WELL': "Well_order r'" -shows "r \o r Osum r'" -using assms Osum_embed Osum_Well_order -unfolding ordLeq_def by blast + assumes FLD: "Field r Int Field r' = {}" and + WELL: "Well_order r" and WELL': "Well_order r'" + shows "r \o r Osum r'" + using assms Osum_embed Osum_Well_order + unfolding ordLeq_def by blast lemma Well_order_embed_copy: -assumes WELL: "well_order_on A r" and - INJ: "inj_on f A" and SUB: "f ` A \ B" -shows "\r'. well_order_on B r' \ r \o r'" + assumes WELL: "well_order_on A r" and + INJ: "inj_on f A" and SUB: "f ` A \ B" + shows "\r'. well_order_on B r' \ r \o r'" proof- have "bij_betw f A (f ` A)" - using INJ inj_on_imp_bij_betw by blast + using INJ inj_on_imp_bij_betw by blast then obtain r'' where "well_order_on (f ` A) r''" and 1: "r =o r''" - using WELL Well_order_iso_copy by blast + using WELL Well_order_iso_copy by blast hence 2: "Well_order r'' \ Field r'' = (f ` A)" - using well_order_on_Well_order by blast - (* *) + using well_order_on_Well_order by blast + (* *) let ?C = "B - (f ` A)" obtain r''' where "well_order_on ?C r'''" - using well_order_on by blast + using well_order_on by blast hence 3: "Well_order r''' \ Field r''' = ?C" - using well_order_on_Well_order by blast - (* *) + using well_order_on_Well_order by blast + (* *) let ?r' = "r'' Osum r'''" have "Field r'' Int Field r''' = {}" - using 2 3 by auto + using 2 3 by auto hence "r'' \o ?r'" using Osum_ordLeq[of r'' r'''] 2 3 by blast hence 4: "r \o ?r'" using 1 ordIso_ordLeq_trans by blast - (* *) + (* *) hence "Well_order ?r'" unfolding ordLeq_def by auto moreover have "Field ?r' = B" using 2 3 SUB by (auto simp add: Field_Osum) @@ -281,174 +185,131 @@ text \The correct phrasing would be ``a maxim of ...", as \\o\ is only a preorder.\ definition isOmax :: "'a rel set \ 'a rel \ bool" -where -"isOmax R r \ r \ R \ (\r' \ R. r' \o r)" + where + "isOmax R r \ r \ R \ (\r' \ R. r' \o r)" definition omax :: "'a rel set \ 'a rel" -where -"omax R == SOME r. isOmax R r" + where + "omax R == SOME r. isOmax R r" lemma exists_isOmax: -assumes "finite R" and "R \ {}" and "\ r \ R. Well_order r" -shows "\ r. isOmax R r" + assumes "finite R" and "R \ {}" and "\ r \ R. Well_order r" + shows "\ r. isOmax R r" proof- have "finite R \ R \ {} \ (\ r \ R. Well_order r) \ (\ r. isOmax R r)" - apply(erule finite_induct) apply(simp add: isOmax_def) + apply(erule finite_induct) apply(simp add: isOmax_def) proof(clarsimp) fix r :: "('a \ 'a) set" and R assume *: "finite R" and **: "r \ R" - and ***: "Well_order r" and ****: "\r\R. Well_order r" - and IH: "R \ {} \ (\p. isOmax R p)" + and ***: "Well_order r" and ****: "\r\R. Well_order r" + and IH: "R \ {} \ (\p. isOmax R p)" let ?R' = "insert r R" show "\p'. (isOmax ?R' p')" proof(cases "R = {}") - assume Case1: "R = {}" - thus ?thesis unfolding isOmax_def using *** - by (simp add: ordLeq_reflexive) + case True + thus ?thesis + by (simp add: "***" isOmax_def ordLeq_reflexive) next - assume Case2: "R \ {}" + case False then obtain p where p: "isOmax R p" using IH by auto - hence 1: "Well_order p" using **** unfolding isOmax_def by simp - {assume Case21: "r \o p" - hence "isOmax ?R' p" using p unfolding isOmax_def by simp - hence ?thesis by auto - } - moreover - {assume Case22: "p \o r" - {fix r' assume "r' \ ?R'" - moreover - {assume "r' \ R" - hence "r' \o p" using p unfolding isOmax_def by simp - hence "r' \o r" using Case22 by(rule ordLeq_transitive) - } - moreover have "r \o r" using *** by(rule ordLeq_reflexive) - ultimately have "r' \o r" by auto - } - hence "isOmax ?R' r" unfolding isOmax_def by simp - hence ?thesis by auto - } - moreover have "r \o p \ p \o r" - using 1 *** ordLeq_total by auto - ultimately show ?thesis by blast + hence "Well_order p" using **** unfolding isOmax_def by simp + then consider "r \o p" | "p \o r" + using *** ordLeq_total by auto + then show ?thesis + proof cases + case 1 + then show ?thesis + using p unfolding isOmax_def by auto + next + case 2 + then show ?thesis + by (metis "***" insert_iff isOmax_def ordLeq_reflexive ordLeq_transitive p) + qed qed qed thus ?thesis using assms by auto qed lemma omax_isOmax: -assumes "finite R" and "R \ {}" and "\ r \ R. Well_order r" -shows "isOmax R (omax R)" -unfolding omax_def using assms -by(simp add: exists_isOmax someI_ex) + assumes "finite R" and "R \ {}" and "\ r \ R. Well_order r" + shows "isOmax R (omax R)" + unfolding omax_def using assms + by(simp add: exists_isOmax someI_ex) lemma omax_in: -assumes "finite R" and "R \ {}" and "\ r \ R. Well_order r" -shows "omax R \ R" -using assms omax_isOmax unfolding isOmax_def by blast + assumes "finite R" and "R \ {}" and "\ r \ R. Well_order r" + shows "omax R \ R" + using assms omax_isOmax unfolding isOmax_def by blast lemma Well_order_omax: -assumes "finite R" and "R \ {}" and "\r\R. Well_order r" -shows "Well_order (omax R)" -using assms apply - apply(drule omax_in) by auto + assumes "finite R" and "R \ {}" and "\r\R. Well_order r" + shows "Well_order (omax R)" + using assms omax_in by blast lemma omax_maxim: -assumes "finite R" and "\ r \ R. Well_order r" and "r \ R" -shows "r \o omax R" -using assms omax_isOmax unfolding isOmax_def by blast + assumes "finite R" and "\ r \ R. Well_order r" and "r \ R" + shows "r \o omax R" + using assms omax_isOmax unfolding isOmax_def by blast lemma omax_ordLeq: -assumes "finite R" and "R \ {}" and *: "\ r \ R. r \o p" -shows "omax R \o p" -proof- - have "\ r \ R. Well_order r" using * unfolding ordLeq_def by simp - thus ?thesis using assms omax_in by auto -qed + assumes "finite R" and "R \ {}" and "\ r \ R. r \o p" + shows "omax R \o p" + by (meson assms omax_in ordLeq_Well_order_simp) lemma omax_ordLess: -assumes "finite R" and "R \ {}" and *: "\ r \ R. r r \ R. Well_order r" using * unfolding ordLess_def by simp - thus ?thesis using assms omax_in by auto -qed + assumes "finite R" and "R \ {}" and "\ r \ R. r r \ R. Well_order r" -and "omax R \o p" and "r \ R" -shows "r \o p" -using assms omax_maxim[of R r] apply simp -using ordLeq_transitive by blast + assumes "finite R" and "\ r \ R. Well_order r" + and "omax R \o p" and "r \ R" + shows "r \o p" + by (meson assms omax_maxim ordLeq_transitive) lemma omax_ordLess_elim: -assumes "finite R" and "\ r \ R. Well_order r" -and "omax R R" -shows "r r \ R. Well_order r" + and "omax R R" + shows "r r \ R. Well_order r" -and "r \ R" and "p \o r" -shows "p \o omax R" -using assms omax_maxim[of R r] apply simp -using ordLeq_transitive by blast + assumes "finite R" and "\ r \ R. Well_order r" + and "r \ R" and "p \o r" + shows "p \o omax R" + by (meson assms omax_maxim ordLeq_transitive) lemma ordLess_omax: -assumes "finite R" and "\ r \ R. Well_order r" -and "r \ R" and "p r \ R. Well_order r" + and "r \ R" and "p {}" and Well_R: "\ r \ R. Well_order r" -and LEQ: "\ p \ P. \ r \ R. p \o r" -shows "omax P \o omax R" -proof- - let ?mp = "omax P" let ?mr = "omax R" - {fix p assume "p \ P" - then obtain r where r: "r \ R" and "p \o r" - using LEQ by blast - moreover have "r <=o ?mr" - using r R Well_R omax_maxim by blast - ultimately have "p <=o ?mr" - using ordLeq_transitive by blast - } - thus "?mp <=o ?mr" - using NE_P P using omax_ordLeq by blast -qed + assumes P: "finite P" and R: "finite R" + and NE_P: "P \ {}" and Well_R: "\ r \ R. Well_order r" + and LEQ: "\ p \ P. \ r \ R. p \o r" + shows "omax P \o omax R" + by (meson LEQ NE_P P R Well_R omax_ordLeq ordLeq_omax) lemma omax_ordLess_mono: -assumes P: "finite P" and R: "finite R" -and NE_P: "P \ {}" and Well_R: "\ r \ R. Well_order r" -and LEQ: "\ p \ P. \ r \ R. p P" - then obtain r where r: "r \ R" and "p {}" and Well_R: "\ r \ R. Well_order r" + and LEQ: "\ p \ P. \ r \ R. p Limit and succesor ordinals\ lemma embed_underS2: -assumes r: "Well_order r" and g: "embed r s g" and a: "a \ Field r" -shows "g ` underS r a = underS s (g a)" + assumes r: "Well_order r" and g: "embed r s g" and a: "a \ Field r" + shows "g ` underS r a = underS s (g a)" by (meson a bij_betw_def embed_underS g r) lemma bij_betw_insert: -assumes "b \ A" and "f b \ A'" and "bij_betw f A A'" -shows "bij_betw f (insert b A) (insert (f b) A')" -using notIn_Un_bij_betw[OF assms] by auto + assumes "b \ A" and "f b \ A'" and "bij_betw f A A'" + shows "bij_betw f (insert b A) (insert (f b) A')" + using notIn_Un_bij_betw[OF assms] by auto context wo_rel begin @@ -459,91 +320,63 @@ by (induct rule: well_order_induct) (rule assms, simp add: underS_def) lemma suc_underS: -assumes B: "B \ Field r" and A: "AboveS B \ {}" and b: "b \ B" -shows "b \ underS (suc B)" -using suc_AboveS[OF B A] b unfolding underS_def AboveS_def by auto + assumes B: "B \ Field r" and A: "AboveS B \ {}" and b: "b \ B" + shows "b \ underS (suc B)" + using suc_AboveS[OF B A] b unfolding underS_def AboveS_def by auto lemma underS_supr: -assumes bA: "b \ underS (supr A)" and A: "A \ Field r" -shows "\ a \ A. b \ underS a" -proof(rule ccontr, auto) + assumes bA: "b \ underS (supr A)" and A: "A \ Field r" + shows "\ a \ A. b \ underS a" +proof(rule ccontr, simp) have bb: "b \ Field r" using bA unfolding underS_def Field_def by auto assume "\a\A. b \ underS a" hence 0: "\a \ A. (a,b) \ r" using A bA unfolding underS_def - by simp (metis REFL in_mono max2_def max2_greater refl_on_domain) - have "(supr A, b) \ r" apply(rule supr_least[OF A bb]) using 0 by auto - thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD) + by simp (metis REFL in_mono max2_def max2_greater refl_on_domain) + have "(supr A, b) \ r" + by (simp add: "0" A bb supr_least) + thus False + by (metis antisymD bA underS_E wo_rel.ANTISYM wo_rel_axioms) qed lemma underS_suc: -assumes bA: "b \ underS (suc A)" and A: "A \ Field r" -shows "\ a \ A. b \ under a" -proof(rule ccontr, auto) + assumes bA: "b \ underS (suc A)" and A: "A \ Field r" + shows "\ a \ A. b \ under a" +proof(rule ccontr, simp) have bb: "b \ Field r" using bA unfolding underS_def Field_def by auto assume "\a\A. b \ under a" - hence 0: "\a \ A. a \ underS b" using A bA unfolding underS_def - by simp (metis (lifting) bb max2_def max2_greater mem_Collect_eq under_def rev_subsetD) - have "(suc A, b) \ r" apply(rule suc_least[OF A bb]) using 0 unfolding underS_def by auto - thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD) + hence 0: "\a \ A. a \ underS b" using A bA + by (metis bb in_mono max2_def max2_greater mem_Collect_eq underS_I under_def) + have "(suc A, b) \ r" + by (metis "0" A bb suc_least underS_E) + thus False + by (metis antisymD bA underS_E wo_rel.ANTISYM wo_rel_axioms) qed lemma (in wo_rel) in_underS_supr: -assumes j: "j \ underS i" and i: "i \ A" and A: "A \ Field r" and AA: "Above A \ {}" -shows "j \ underS (supr A)" -proof- - have "(i,supr A) \ r" using supr_greater[OF A AA i] . - thus ?thesis using j unfolding underS_def - by simp (metis REFL TRANS max2_def max2_equals1 refl_on_domain transD) -qed + assumes "j \ underS i" and "i \ A" and "A \ Field r" and "Above A \ {}" + shows "j \ underS (supr A)" + by (meson assms LIN in_mono supr_greater supr_inField underS_incl_iff) lemma inj_on_Field: -assumes A: "A \ Field r" and f: "\ a b. \a \ A; b \ A; a \ underS b\ \ f a \ f b" -shows "inj_on f A" -unfolding inj_on_def proof safe - fix a b assume a: "a \ A" and b: "b \ A" and fab: "f a = f b" - {assume "a \ underS b" - hence False using f[OF a b] fab by auto - } - moreover - {assume "b \ underS a" - hence False using f[OF b a] fab by auto - } - ultimately show "a = b" using TOTALS A a b unfolding underS_def by auto -qed + assumes A: "A \ Field r" and f: "\ a b. \a \ A; b \ A; a \ underS b\ \ f a \ f b" + shows "inj_on f A" + by (smt (verit) A f in_notinI inj_on_def subsetD underS_I) lemma ofilter_init_seg_of: -assumes "ofilter F" -shows "Restr r F initial_segment_of r" -using assms unfolding ofilter_def init_seg_of_def under_def by auto + assumes "ofilter F" + shows "Restr r F initial_segment_of r" + using assms unfolding ofilter_def init_seg_of_def under_def by auto lemma underS_init_seg_of_Collect: -assumes "\j1 j2. \j2 \ underS i; (j1, j2) \ r\ \ R j1 initial_segment_of R j2" -shows "{R j |j. j \ underS i} \ Chains init_seg_of" -unfolding Chains_def proof safe - fix j ja assume jS: "j \ underS i" and jaS: "ja \ underS i" - and init: "(R ja, R j) \ init_seg_of" - hence jja: "{j,ja} \ Field r" and j: "j \ Field r" and ja: "ja \ Field r" - and jjai: "(j,i) \ r" "(ja,i) \ r" - and i: "i \ {j,ja}" unfolding Field_def underS_def by auto - have jj: "(j,j) \ r" and jaja: "(ja,ja) \ r" using j ja by (metis in_notinI)+ - show "R j initial_segment_of R ja" - using jja init jjai i - by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: underS_def) -qed + assumes "\j1 j2. \j2 \ underS i; (j1, j2) \ r\ \ R j1 initial_segment_of R j2" + shows "{R j |j. j \ underS i} \ Chains init_seg_of" + using TOTALS assms + by (clarsimp simp: Chains_def) (meson BNF_Least_Fixpoint.underS_Field) lemma (in wo_rel) Field_init_seg_of_Collect: -assumes "\j1 j2. \j2 \ Field r; (j1, j2) \ r\ \ R j1 initial_segment_of R j2" -shows "{R j |j. j \ Field r} \ Chains init_seg_of" -unfolding Chains_def proof safe - fix j ja assume jS: "j \ Field r" and jaS: "ja \ Field r" - and init: "(R ja, R j) \ init_seg_of" - hence jja: "{j,ja} \ Field r" and j: "j \ Field r" and ja: "ja \ Field r" - unfolding Field_def underS_def by auto - have jj: "(j,j) \ r" and jaja: "(ja,ja) \ r" using j ja by (metis in_notinI)+ - show "R j initial_segment_of R ja" - using jja init - by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: Field_def) -qed + assumes "\j1 j2. \j2 \ Field r; (j1, j2) \ r\ \ R j1 initial_segment_of R j2" + shows "{R j |j. j \ Field r} \ Chains init_seg_of" + using TOTALS assms by (auto simp: Chains_def) subsubsection \Successor and limit elements of an ordinal\ @@ -556,77 +389,63 @@ definition "isLim i \ \ isSucc i" lemma zero_smallest[simp]: -assumes "j \ Field r" shows "(zero, j) \ r" -unfolding zero_def -by (metis AboveS_Field assms subset_AboveS_UnderS subset_antisym subset_refl suc_def suc_least_AboveS) + assumes "j \ Field r" shows "(zero, j) \ r" + by (simp add: assms wo_rel.ofilter_linord wo_rel_axioms zero_def) lemma zero_in_Field: assumes "Field r \ {}" shows "zero \ Field r" -using assms unfolding zero_def by (metis Field_ofilter minim_in ofilter_def) + using assms unfolding zero_def by (metis Field_ofilter minim_in ofilter_def) lemma leq_zero_imp[simp]: -"(x, zero) \ r \ x = zero" -by (metis ANTISYM WELL antisymD well_order_on_domain zero_smallest) + "(x, zero) \ r \ x = zero" + by (metis ANTISYM WELL antisymD well_order_on_domain zero_smallest) lemma leq_zero[simp]: -assumes "Field r \ {}" shows "(x, zero) \ r \ x = zero" -using zero_in_Field[OF assms] in_notinI[of x zero] by auto + assumes "Field r \ {}" shows "(x, zero) \ r \ x = zero" + using zero_in_Field[OF assms] in_notinI[of x zero] by auto lemma under_zero[simp]: -assumes "Field r \ {}" shows "under zero = {zero}" -using assms unfolding under_def by auto + assumes "Field r \ {}" shows "under zero = {zero}" + using assms unfolding under_def by auto lemma underS_zero[simp,intro]: "underS zero = {}" -unfolding underS_def by auto + unfolding underS_def by auto lemma isSucc_succ: "aboveS i \ {} \ isSucc (succ i)" -unfolding isSucc_def succ_def by auto + unfolding isSucc_def succ_def by auto lemma succ_in_diff: -assumes "aboveS i \ {}" shows "(i,succ i) \ r \ succ i \ i" -using assms suc_greater[of "{i}"] unfolding succ_def AboveS_def aboveS_def Field_def by auto + assumes "aboveS i \ {}" shows "(i,succ i) \ r \ succ i \ i" + using assms suc_greater[of "{i}"] unfolding succ_def AboveS_def aboveS_def Field_def by auto lemmas succ_in[simp] = succ_in_diff[THEN conjunct1] lemmas succ_diff[simp] = succ_in_diff[THEN conjunct2] lemma succ_in_Field[simp]: -assumes "aboveS i \ {}" shows "succ i \ Field r" -using succ_in[OF assms] unfolding Field_def by auto + assumes "aboveS i \ {}" shows "succ i \ Field r" + using succ_in[OF assms] unfolding Field_def by auto lemma succ_not_in: -assumes "aboveS i \ {}" shows "(succ i, i) \ r" -proof - assume 1: "(succ i, i) \ r" - hence "succ i \ Field r \ i \ Field r" unfolding Field_def by auto - hence "(i, succ i) \ r \ succ i \ i" using assms by auto - thus False using 1 by (metis ANTISYM antisymD) -qed + assumes "aboveS i \ {}" shows "(succ i, i) \ r" + by (metis FieldI2 assms max2_equals1 max2_equals2 succ_diff succ_in) lemma not_isSucc_zero: "\ isSucc zero" -proof - assume *: "isSucc zero" - then obtain i where "aboveS i \ {}" and 1: "minim (Field r) = succ i" - unfolding isSucc_def zero_def by auto - hence "succ i \ Field r" by auto - with * show False - by (metis REFL isSucc_def minim_least refl_on_domain - subset_refl succ_in succ_not_in zero_def) -qed + by (metis isSucc_def leq_zero_imp succ_in_diff) lemma isLim_zero[simp]: "isLim zero" by (metis isLim_def not_isSucc_zero) lemma succ_smallest: -assumes "(i,j) \ r" and "i \ j" -shows "(succ i, j) \ r" -unfolding succ_def apply(rule suc_least) -using assms unfolding Field_def by auto + assumes "(i,j) \ r" and "i \ j" + shows "(succ i, j) \ r" + by (metis Field_iff assms empty_subsetI insert_subset singletonD suc_least succ_def) lemma isLim_supr: -assumes f: "i \ Field r" and l: "isLim i" -shows "i = supr (underS i)" + assumes f: "i \ Field r" and l: "isLim i" + shows "i = supr (underS i)" proof(rule equals_supr) fix j assume j: "j \ Field r" and 1: "\ j'. j' \ underS i \ (j',j) \ r" - show "(i,j) \ r" proof(intro in_notinI[OF _ f j], safe) + show "(i,j) \ r" + proof(intro in_notinI[OF _ f j], safe) assume ji: "(j,i) \ r" "j \ i" hence a: "aboveS j \ {}" unfolding aboveS_def by auto hence "i \ succ j" using l unfolding isLim_def isSucc_def by auto @@ -635,17 +454,19 @@ hence "(succ j, j) \ r" using 1 by auto thus False using succ_not_in[OF a] by simp qed -qed(insert f, unfold underS_def Field_def, auto) +qed(use f underS_def Field_def in fastforce)+ definition "pred i \ SOME j. j \ Field r \ aboveS j \ {} \ succ j = i" lemma pred_Field_succ: -assumes "isSucc i" shows "pred i \ Field r \ aboveS (pred i) \ {} \ succ (pred i) = i" + assumes "isSucc i" shows "pred i \ Field r \ aboveS (pred i) \ {} \ succ (pred i) = i" proof- - obtain j where a: "aboveS j \ {}" and i: "i = succ j" using assms unfolding isSucc_def by auto - have 1: "j \ Field r" "j \ i" unfolding Field_def i - using succ_diff[OF a] a unfolding aboveS_def by auto - show ?thesis unfolding pred_def apply(rule someI_ex) using 1 i a by auto + obtain j where j: "aboveS j \ {}" "i = succ j" + using assms unfolding isSucc_def by auto + then obtain "j \ Field r" "j \ i" + by (metis FieldI1 succ_diff succ_in) + then show ?thesis unfolding pred_def + by (metis (mono_tags, lifting) j someI_ex) qed lemmas pred_Field[simp] = pred_Field_succ[THEN conjunct1] @@ -653,165 +474,100 @@ lemmas succ_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct2] lemma isSucc_pred_in: -assumes "isSucc i" shows "(pred i, i) \ r" -proof- - define j where "j = pred i" - have i: "i = succ j" using assms unfolding j_def by simp - have a: "aboveS j \ {}" unfolding j_def using assms by auto - show ?thesis unfolding j_def[symmetric] unfolding i using succ_in[OF a] . -qed + assumes "isSucc i" shows "(pred i, i) \ r" + by (metis assms pred_Field_succ succ_in) lemma isSucc_pred_diff: -assumes "isSucc i" shows "pred i \ i" -by (metis aboveS_pred assms succ_diff succ_pred) + assumes "isSucc i" shows "pred i \ i" + by (metis aboveS_pred assms succ_diff succ_pred) (* todo: pred maximal, pred injective? *) lemma succ_inj[simp]: -assumes "aboveS i \ {}" and "aboveS j \ {}" -shows "succ i = succ j \ i = j" -proof safe - assume s: "succ i = succ j" - {assume "i \ j" and "(i,j) \ r" - hence "(succ i, j) \ r" using assms by (metis succ_smallest) - hence False using s assms by (metis succ_not_in) - } - moreover - {assume "i \ j" and "(j,i) \ r" - hence "(succ j, i) \ r" using assms by (metis succ_smallest) - hence False using s assms by (metis succ_not_in) - } - ultimately show "i = j" by (metis TOTALS WELL assms(1) assms(2) succ_in_diff well_order_on_domain) -qed + assumes "aboveS i \ {}" and "aboveS j \ {}" + shows "succ i = succ j \ i = j" + by (metis FieldI1 assms succ_def succ_in supr_under under_underS_suc) lemma pred_succ[simp]: -assumes "aboveS j \ {}" shows "pred (succ j) = j" -unfolding pred_def apply(rule some_equality) -using assms apply(force simp: Field_def aboveS_def) -by (metis assms succ_inj) + assumes "aboveS j \ {}" shows "pred (succ j) = j" + using assms isSucc_def pred_Field_succ succ_inj by blast lemma less_succ[simp]: -assumes "aboveS i \ {}" -shows "(j, succ i) \ r \ (j,i) \ r \ j = succ i" -apply safe - apply (metis WELL assms in_notinI well_order_on_domain suc_singl_pred succ_def succ_in_diff) - apply (metis (opaque_lifting, full_types) REFL TRANS assms max2_def max2_equals1 refl_on_domain succ_in_Field succ_not_in transD) - apply (metis assms in_notinI succ_in_Field) -done + assumes "aboveS i \ {}" + shows "(j, succ i) \ r \ (j,i) \ r \ j = succ i" + by (metis FieldI1 assms in_notinI max2_equals1 max2_equals2 max2_iff succ_in succ_smallest) lemma underS_succ[simp]: -assumes "aboveS i \ {}" -shows "underS (succ i) = under i" -unfolding underS_def under_def by (auto simp: assms succ_not_in) + assumes "aboveS i \ {}" + shows "underS (succ i) = under i" + unfolding underS_def under_def by (auto simp: assms succ_not_in) lemma succ_mono: -assumes "aboveS j \ {}" and "(i,j) \ r" -shows "(succ i, succ j) \ r" -by (metis (full_types) assms less_succ succ_smallest) + assumes "aboveS j \ {}" and "(i,j) \ r" + shows "(succ i, succ j) \ r" + by (metis (full_types) assms less_succ succ_smallest) lemma under_succ[simp]: -assumes "aboveS i \ {}" -shows "under (succ i) = insert (succ i) (under i)" -using less_succ[OF assms] unfolding under_def by auto + assumes "aboveS i \ {}" + shows "under (succ i) = insert (succ i) (under i)" + using less_succ[OF assms] unfolding under_def by auto definition mergeSL :: "('a \ 'b \ 'b) \ (('a \ 'b) \ 'a \ 'b) \ ('a \ 'b) \ 'a \ 'b" -where -"mergeSL S L f i \ - if isSucc i then S (pred i) (f (pred i)) - else L f i" + where + "mergeSL S L f i \ if isSucc i then S (pred i) (f (pred i)) else L f i" subsubsection \Well-order recursion with (zero), succesor, and limit\ definition worecSL :: "('a \ 'b \ 'b) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" -where "worecSL S L \ worec (mergeSL S L)" + where "worecSL S L \ worec (mergeSL S L)" definition "adm_woL L \ \f g i. isLim i \ (\j\underS i. f j = g j) \ L f i = L g i" -lemma mergeSL: -assumes "adm_woL L" shows "adm_wo (mergeSL S L)" -unfolding adm_wo_def proof safe - fix f g :: "'a => 'b" and i :: 'a - assume 1: "\j\underS i. f j = g j" - show "mergeSL S L f i = mergeSL S L g i" - proof(cases "isSucc i") - case True - hence "pred i \ underS i" unfolding underS_def using isSucc_pred_in isSucc_pred_diff by auto - thus ?thesis using True 1 unfolding mergeSL_def by auto - next - case False hence "isLim i" unfolding isLim_def by auto - thus ?thesis using assms False 1 unfolding mergeSL_def adm_woL_def by auto - qed -qed +lemma mergeSL: "adm_woL L \adm_wo (mergeSL S L)" + unfolding adm_wo_def adm_woL_def isLim_def + by (smt (verit, ccfv_threshold) isSucc_pred_diff isSucc_pred_in mergeSL_def underS_I) lemma worec_fixpoint1: "adm_wo H \ worec H i = H (worec H) i" -by (metis worec_fixpoint) + by (metis worec_fixpoint) lemma worecSL_isSucc: -assumes a: "adm_woL L" and i: "isSucc i" -shows "worecSL S L i = S (pred i) (worecSL S L (pred i))" -proof- - let ?H = "mergeSL S L" - have "worecSL S L i = ?H (worec ?H) i" - unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] . - also have "... = S (pred i) (worecSL S L (pred i))" - unfolding worecSL_def mergeSL_def using i by simp - finally show ?thesis . -qed + assumes a: "adm_woL L" and i: "isSucc i" + shows "worecSL S L i = S (pred i) (worecSL S L (pred i))" + by (metis a i mergeSL mergeSL_def worecSL_def worec_fixpoint) lemma worecSL_succ: -assumes a: "adm_woL L" and i: "aboveS j \ {}" -shows "worecSL S L (succ j) = S j (worecSL S L j)" -proof- - define i where "i = succ j" - have i: "isSucc i" by (metis i i_def isSucc_succ) - have ij: "j = pred i" unfolding i_def using assms by simp - have 0: "succ (pred i) = i" using i by simp - show ?thesis unfolding ij using worecSL_isSucc[OF a i] unfolding 0 . -qed + assumes a: "adm_woL L" and i: "aboveS j \ {}" + shows "worecSL S L (succ j) = S j (worecSL S L j)" + by (simp add: a i isSucc_succ worecSL_isSucc) lemma worecSL_isLim: -assumes a: "adm_woL L" and i: "isLim i" -shows "worecSL S L i = L (worecSL S L) i" -proof- - let ?H = "mergeSL S L" - have "worecSL S L i = ?H (worec ?H) i" - unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] . - also have "... = L (worecSL S L) i" - using i unfolding worecSL_def mergeSL_def isLim_def by simp - finally show ?thesis . -qed + assumes a: "adm_woL L" and i: "isLim i" + shows "worecSL S L i = L (worecSL S L) i" + by (metis a i isLim_def mergeSL mergeSL_def worecSL_def worec_fixpoint) definition worecZSL :: "'b \ ('a \ 'b \ 'b) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" -where "worecZSL Z S L \ worecSL S (\ f a. if a = zero then Z else L f a)" + where "worecZSL Z S L \ worecSL S (\ f a. if a = zero then Z else L f a)" lemma worecZSL_zero: -assumes a: "adm_woL L" -shows "worecZSL Z S L zero = Z" -proof- - let ?L = "\ f a. if a = zero then Z else L f a" - have "worecZSL Z S L zero = ?L (worecZSL Z S L) zero" - unfolding worecZSL_def apply(rule worecSL_isLim) - using assms unfolding adm_woL_def by auto - also have "... = Z" by simp - finally show ?thesis . -qed + assumes a: "adm_woL L" + shows "worecZSL Z S L zero = Z" + by (smt (verit, best) adm_woL_def assms isLim_zero worecSL_isLim worecZSL_def) lemma worecZSL_succ: -assumes a: "adm_woL L" and i: "aboveS j \ {}" -shows "worecZSL Z S L (succ j) = S j (worecZSL Z S L j)" -unfolding worecZSL_def apply(rule worecSL_succ) -using assms unfolding adm_woL_def by auto + assumes a: "adm_woL L" and i: "aboveS j \ {}" + shows "worecZSL Z S L (succ j) = S j (worecZSL Z S L j)" + unfolding worecZSL_def + by (smt (verit, best) a adm_woL_def i worecSL_succ) lemma worecZSL_isLim: -assumes a: "adm_woL L" and "isLim i" and "i \ zero" -shows "worecZSL Z S L i = L (worecZSL Z S L) i" + assumes a: "adm_woL L" and "isLim i" and "i \ zero" + shows "worecZSL Z S L i = L (worecZSL Z S L) i" proof- let ?L = "\ f a. if a = zero then Z else L f a" have "worecZSL Z S L i = ?L (worecZSL Z S L) i" - unfolding worecZSL_def apply(rule worecSL_isLim) - using assms unfolding adm_woL_def by auto - also have "... = L (worecZSL Z S L) i" using assms by simp + unfolding worecZSL_def by (smt (verit, best) adm_woL_def assms worecSL_isLim) + also have "\ = L (worecZSL Z S L) i" using assms by simp finally show ?thesis . qed @@ -819,70 +575,54 @@ subsubsection \Well-order succ-lim induction\ lemma ord_cases: -obtains j where "i = succ j" and "aboveS j \ {}" | "isLim i" -by (metis isLim_def isSucc_def) + obtains j where "i = succ j" and "aboveS j \ {}" | "isLim i" + by (metis isLim_def isSucc_def) lemma well_order_inductSL[case_names Suc Lim]: -assumes SUCC: "\i. \aboveS i \ {}; P i\ \ P (succ i)" and -LIM: "\i. \isLim i; \j. j \ underS i \ P j\ \ P i" -shows "P i" + assumes "\i. \aboveS i \ {}; P i\ \ P (succ i)" "\i. \isLim i; \j. j \ underS i \ P j\ \ P i" + shows "P i" proof(induction rule: well_order_induct) - fix i assume 0: "\j. j \ i \ (j, i) \ r \ P j" - show "P i" proof(cases i rule: ord_cases) - fix j assume i: "i = succ j" and j: "aboveS j \ {}" - hence "j \ i \ (j, i) \ r" by (metis succ_diff succ_in) - hence 1: "P j" using 0 by simp - show "P i" unfolding i apply(rule SUCC) using 1 j by auto - qed(insert 0 LIM, unfold underS_def, auto) + case (1 x) + then show ?case + by (metis assms ord_cases succ_diff succ_in underS_E) qed lemma well_order_inductZSL[case_names Zero Suc Lim]: -assumes ZERO: "P zero" -and SUCC: "\i. \aboveS i \ {}; P i\ \ P (succ i)" and -LIM: "\i. \isLim i; i \ zero; \j. j \ underS i \ P j\ \ P i" -shows "P i" -apply(induction rule: well_order_inductSL) using assms by auto + assumes "P zero" + and "\i. \aboveS i \ {}; P i\ \ P (succ i)" and + "\i. \isLim i; i \ zero; \j. j \ underS i \ P j\ \ P i" + shows "P i" + by (metis assms well_order_inductSL) (* Succesor and limit ordinals *) definition "isSuccOrd \ \ j \ Field r. \ i \ Field r. (i,j) \ r" definition "isLimOrd \ \ isSuccOrd" lemma isLimOrd_succ: -assumes isLimOrd and "i \ Field r" -shows "succ i \ Field r" -using assms unfolding isLimOrd_def isSuccOrd_def -by (metis REFL in_notinI refl_on_domain succ_smallest) + assumes isLimOrd and "i \ Field r" + shows "succ i \ Field r" + using assms unfolding isLimOrd_def isSuccOrd_def + by (metis REFL in_notinI refl_on_domain succ_smallest) lemma isLimOrd_aboveS: -assumes l: isLimOrd and i: "i \ Field r" -shows "aboveS i \ {}" + assumes l: isLimOrd and i: "i \ Field r" + shows "aboveS i \ {}" proof- obtain j where "j \ Field r" and "(j,i) \ r" - using assms unfolding isLimOrd_def isSuccOrd_def by auto + using assms unfolding isLimOrd_def isSuccOrd_def by auto hence "(i,j) \ r \ j \ i" by (metis i max2_def max2_greater) thus ?thesis unfolding aboveS_def by auto qed lemma succ_aboveS_isLimOrd: -assumes "\ i \ Field r. aboveS i \ {} \ succ i \ Field r" -shows isLimOrd -unfolding isLimOrd_def isSuccOrd_def proof safe - fix j assume j: "j \ Field r" "\i\Field r. (i, j) \ r" - hence "(succ j, j) \ r" using assms by auto - moreover have "aboveS j \ {}" using assms j unfolding aboveS_def by auto - ultimately show False by (metis succ_not_in) -qed + assumes "\ i \ Field r. aboveS i \ {} \ succ i \ Field r" + shows isLimOrd + using assms isLimOrd_def isSuccOrd_def succ_not_in by blast lemma isLim_iff: -assumes l: "isLim i" and j: "j \ underS i" -shows "\ k. k \ underS i \ j \ underS k" -proof- - have a: "aboveS j \ {}" using j unfolding underS_def aboveS_def by auto - show ?thesis apply(rule exI[of _ "succ j"]) apply safe - using assms a unfolding underS_def isLim_def - apply (metis (lifting, full_types) isSucc_def mem_Collect_eq succ_smallest) - by (metis (lifting, full_types) a mem_Collect_eq succ_diff succ_in) -qed + assumes l: "isLim i" and j: "j \ underS i" + shows "\ k. k \ underS i \ j \ underS k" + by (metis Order_Relation.underS_Field empty_iff isLim_supr j l underS_empty underS_supr) end (* context wo_rel *) @@ -903,124 +643,73 @@ definition "oproj r s f \ Field s \ f ` (Field r) \ compat r s f" lemma oproj_in: -assumes "oproj r s f" and "(a,a') \ r" -shows "(f a, f a') \ s" -using assms unfolding oproj_def compat_def by auto + assumes "oproj r s f" and "(a,a') \ r" + shows "(f a, f a') \ s" + using assms unfolding oproj_def compat_def by auto lemma oproj_Field: -assumes f: "oproj r s f" and a: "a \ Field r" -shows "f a \ Field s" -using oproj_in[OF f] a unfolding Field_def by auto + assumes f: "oproj r s f" and a: "a \ Field r" + shows "f a \ Field s" + using oproj_in[OF f] a unfolding Field_def by auto lemma oproj_Field2: -assumes f: "oproj r s f" and a: "b \ Field s" -shows "\ a \ Field r. f a = b" -using assms unfolding oproj_def by auto + assumes f: "oproj r s f" and a: "b \ Field s" + shows "\ a \ Field r. f a = b" + using assms unfolding oproj_def by auto lemma oproj_under: -assumes f: "oproj r s f" and a: "a \ under r a'" -shows "f a \ under s (f a')" -using oproj_in[OF f] a unfolding under_def by auto + assumes f: "oproj r s f" and a: "a \ under r a'" + shows "f a \ under s (f a')" + using oproj_in[OF f] a unfolding under_def by auto (* An ordinal is embedded in another whenever it is embedded as an order (not necessarily as initial segment):*) theorem embedI: -assumes r: "Well_order r" and s: "Well_order s" -and f: "\ a. a \ Field r \ f a \ Field s \ f ` underS r a \ underS s (f a)" -shows "\ g. embed r s g" + assumes r: "Well_order r" and s: "Well_order s" + and f: "\ a. a \ Field r \ f a \ Field s \ f ` underS r a \ underS s (f a)" + shows "\ g. embed r s g" proof- interpret r: wo_rel r by unfold_locales (rule r) interpret s: wo_rel s by unfold_locales (rule s) let ?G = "\ g a. suc s (g ` underS r a)" define g where "g = worec r ?G" have adm: "adm_wo r ?G" unfolding r.adm_wo_def image_def by auto - (* *) + (* *) {fix a assume "a \ Field r" - hence "bij_betw g (under r a) (under s (g a)) \ + hence "bij_betw g (under r a) (under s (g a)) \ g a \ under s (f a)" - proof(induction a rule: r.underS_induct) - case (1 a) - hence a: "a \ Field r" - and IH1a: "\ a1. a1 \ underS r a \ inj_on g (under r a1)" - and IH1b: "\ a1. a1 \ underS r a \ g ` under r a1 = under s (g a1)" - and IH2: "\ a1. a1 \ underS r a \ g a1 \ under s (f a1)" - unfolding underS_def Field_def bij_betw_def by auto - have fa: "f a \ Field s" using f[OF a] by auto - have g: "g a = suc s (g ` underS r a)" - using r.worec_fixpoint[OF adm] unfolding g_def fun_eq_iff by blast - have A0: "g ` underS r a \ Field s" - using IH1b by (metis IH2 image_subsetI in_mono under_Field) - {fix a1 assume a1: "a1 \ underS r a" - from IH2[OF this] have "g a1 \ under s (f a1)" . - moreover have "f a1 \ underS s (f a)" using f[OF a] a1 by auto - ultimately have "g a1 \ underS s (f a)" by (metis s.ANTISYM s.TRANS under_underS_trans) - } - hence "f a \ AboveS s (g ` underS r a)" unfolding AboveS_def - using fa by simp (metis (lifting, full_types) mem_Collect_eq underS_def) - hence A: "AboveS s (g ` underS r a) \ {}" by auto - have B: "\ a1. a1 \ underS r a \ g a1 \ underS s (g a)" - unfolding g apply(rule s.suc_underS[OF A0 A]) by auto - {fix a1 a2 assume a2: "a2 \ underS r a" and 1: "a1 \ underS r a2" - hence a12: "{a1,a2} \ under r a2" and "a1 \ a2" using r.REFL a - unfolding underS_def under_def refl_on_def Field_def by auto - hence "g a1 \ g a2" using IH1a[OF a2] unfolding inj_on_def by auto - hence "g a1 \ underS s (g a2)" using IH1b[OF a2] a12 - unfolding underS_def under_def by auto - } note C = this - have ga: "g a \ Field s" unfolding g using s.suc_inField[OF A0 A] . - have aa: "a \ under r a" - using a r.REFL unfolding under_def underS_def refl_on_def by auto - show ?case proof safe - show "bij_betw g (under r a) (under s (g a))" unfolding bij_betw_def proof safe - show "inj_on g (under r a)" proof(rule r.inj_on_Field) - fix a1 a2 assume "a1 \ under r a" and a2: "a2 \ under r a" and a1: "a1 \ underS r a2" - hence a22: "a2 \ under r a2" and a12: "a1 \ under r a2" "a1 \ a2" - using a r.REFL unfolding under_def underS_def refl_on_def by auto - show "g a1 \ g a2" - proof(cases "a2 = a") - case False hence "a2 \ underS r a" - using a2 unfolding underS_def under_def by auto - from IH1a[OF this] show ?thesis using a12 a22 unfolding inj_on_def by auto - qed(insert B a1, unfold underS_def, auto) - qed(unfold under_def Field_def, auto) - next - fix a1 assume a1: "a1 \ under r a" - show "g a1 \ under s (g a)" - proof(cases "a1 = a") - case True thus ?thesis - using ga s.REFL unfolding refl_on_def under_def by auto - next - case False - hence a1: "a1 \ underS r a" using a1 unfolding underS_def under_def by auto - thus ?thesis using B unfolding underS_def under_def by auto - qed - next - fix b1 assume b1: "b1 \ under s (g a)" - show "b1 \ g ` under r a" - proof(cases "b1 = g a") - case True thus ?thesis using aa by auto - next - case False - hence "b1 \ underS s (g a)" using b1 unfolding underS_def under_def by auto - from s.underS_suc[OF this[unfolded g] A0] - obtain a1 where a1: "a1 \ underS r a" and b1: "b1 \ under s (g a1)" by auto - obtain a2 where "a2 \ under r a1" and b1: "b1 = g a2" using IH1b[OF a1] b1 by auto - hence "a2 \ under r a" using a1 - by (metis r.ANTISYM r.TRANS in_mono underS_subset_under under_underS_trans) - thus ?thesis using b1 by auto - qed - qed - next - have "(g a, f a) \ s" unfolding g proof(rule s.suc_least[OF A0]) - fix b1 assume "b1 \ g ` underS r a" - then obtain a1 where a1: "b1 = g a1" and a1: "a1 \ underS r a" by auto - hence "b1 \ underS s (f a)" - using a by (metis \\a1. a1 \ underS r a \ g a1 \ underS s (f a)\) - thus "f a \ b1 \ (b1, f a) \ s" unfolding underS_def by auto - qed(insert fa, auto) - thus "g a \ under s (f a)" unfolding under_def by auto - qed - qed + proof(induction a rule: r.underS_induct) + case (1 a) + hence a: "a \ Field r" + and IH1a: "\ a1. a1 \ underS r a \ inj_on g (under r a1)" + and IH1b: "\ a1. a1 \ underS r a \ g ` under r a1 = under s (g a1)" + and IH2: "\ a1. a1 \ underS r a \ g a1 \ under s (f a1)" + unfolding underS_def Field_def bij_betw_def by auto + have fa: "f a \ Field s" using f[OF a] by auto + have g: "g a = suc s (g ` underS r a)" + using r.worec_fixpoint[OF adm] unfolding g_def fun_eq_iff by blast + have A0: "g ` underS r a \ Field s" + using IH1b by (metis IH2 image_subsetI in_mono under_Field) + {fix a1 assume a1: "a1 \ underS r a" + from IH2[OF this] have "g a1 \ under s (f a1)" . + moreover have "f a1 \ underS s (f a)" using f[OF a] a1 by auto + ultimately have "g a1 \ underS s (f a)" by (metis s.ANTISYM s.TRANS under_underS_trans) + } + hence fa_in: "f a \ AboveS s (g ` underS r a)" unfolding AboveS_def + using fa by simp (metis (lifting, full_types) mem_Collect_eq underS_def) + hence A: "AboveS s (g ` underS r a) \ {}" by auto + have ga: "g a \ Field s" unfolding g using s.suc_inField[OF A0 A] . + show ?case + unfolding bij_betw_def + proof (intro conjI) + show "inj_on g (r.under a)" + by (metis A IH1a IH1b a bij_betw_def g ga r s s.suc_greater subsetI wellorders_totally_ordered_aux) + show "g ` r.under a = s.under (g a)" + by (metis A A0 IH1a IH1b a bij_betw_def g ga r s s.suc_greater wellorders_totally_ordered_aux) + show "g a \ s.under (f a)" + by (simp add: fa_in g s.suc_least_AboveS under_def) + qed + qed } thus ?thesis unfolding embed_def by auto qed @@ -1028,39 +717,38 @@ corollary ordLeq_def2: "r \o s \ Well_order r \ Well_order s \ (\ f. \ a \ Field r. f a \ Field s \ f ` underS r a \ underS s (f a))" -using embed_in_Field[of r s] embed_underS2[of r s] embedI[of r s] -unfolding ordLeq_def by fast + using embed_in_Field[of r s] embed_underS2[of r s] embedI[of r s] + unfolding ordLeq_def by fast lemma iso_oproj: assumes r: "Well_order r" and s: "Well_order s" and f: "iso r s f" shows "oproj r s f" -using assms unfolding oproj_def -by (subst (asm) iso_iff3) (auto simp: bij_betw_def) + by (metis embed_Field f iso_Field iso_iff iso_iff3 oproj_def r s) theorem oproj_embed: -assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f" -shows "\ g. embed s r g" + assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f" + shows "\ g. embed s r g" proof (rule embedI[OF s r, of "inv_into (Field r) f"], unfold underS_def, safe) fix b assume "b \ Field s" - thus "inv_into (Field r) f b \ Field r" using oproj_Field2[OF f] by (metis imageI inv_into_into) + thus "inv_into (Field r) f b \ Field r" + using oproj_Field2[OF f] by (metis imageI inv_into_into) next fix a b assume "b \ Field s" "a \ b" "(a, b) \ s" "inv_into (Field r) f a = inv_into (Field r) f b" - with f show False by (auto dest!: inv_into_injective simp: Field_def oproj_def) + with f show False + by (meson FieldI1 in_mono inv_into_injective oproj_def) next fix a b assume *: "b \ Field s" "a \ b" "(a, b) \ s" - { assume "(inv_into (Field r) f a, inv_into (Field r) f b) \ r" + { assume notin: "(inv_into (Field r) f a, inv_into (Field r) f b) \ r" moreover from *(3) have "a \ Field s" unfolding Field_def by auto - with *(1,2) have - "inv_into (Field r) f a \ Field r" "inv_into (Field r) f b \ Field r" - "inv_into (Field r) f a \ inv_into (Field r) f b" - by (auto dest!: oproj_Field2[OF f] inv_into_injective intro!: inv_into_into) + then have "(inv_into (Field r) f b, inv_into (Field r) f a) \ r" + by (meson "*"(1) notin f in_mono inv_into_into oproj_def r wo_rel.in_notinI wo_rel.intro) ultimately have "(inv_into (Field r) f b, inv_into (Field r) f a) \ r" using r by (auto simp: well_order_on_def linear_order_on_def total_on_def) with f[unfolded oproj_def compat_def] *(1) \a \ Field s\ f_inv_into_f[of b f "Field r"] f_inv_into_f[of a f "Field r"] - have "(b, a) \ s" by (metis in_mono) + have "(b, a) \ s" by (metis in_mono) with *(2,3) s have False by (auto simp: well_order_on_def linear_order_on_def partial_order_on_def antisym_def) } thus "(inv_into (Field r) f a, inv_into (Field r) f b) \ r" by blast @@ -1069,6 +757,6 @@ corollary oproj_ordLeq: assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f" shows "s \o r" - using oproj_embed[OF assms] r s unfolding ordLeq_def by auto + using f oproj_embed ordLess_iff ordLess_or_ordLeq r s by blast end diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Cardinals/Wellorder_Embedding.thy --- a/src/HOL/Cardinals/Wellorder_Embedding.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Cardinals/Wellorder_Embedding.thy Sun Jan 15 20:00:44 2023 +0100 @@ -8,22 +8,22 @@ section \Well-Order Embeddings\ theory Wellorder_Embedding -imports Fun_More Wellorder_Relation + imports Fun_More Wellorder_Relation begin subsection \Auxiliaries\ lemma UNION_bij_betw_ofilter: -assumes WELL: "Well_order r" and - OF: "\ i. i \ I \ ofilter r (A i)" and - BIJ: "\ i. i \ I \ bij_betw f (A i) (A' i)" -shows "bij_betw f (\i \ I. A i) (\i \ I. A' i)" + assumes WELL: "Well_order r" and + OF: "\ i. i \ I \ ofilter r (A i)" and + BIJ: "\ i. i \ I \ bij_betw f (A i) (A' i)" + shows "bij_betw f (\i \ I. A i) (\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 + using wo_rel.ofilter_linord[of r] OF by blast with WELL BIJ show ?thesis - by (auto simp add: bij_betw_UNION_chain) + by (auto simp add: bij_betw_UNION_chain) qed @@ -31,144 +31,98 @@ functions\ lemma embed_halfcong: -assumes EQ: "\ a. a \ Field r \ f a = g a" and - EMB: "embed r r' f" -shows "embed r r' g" -proof(unfold embed_def, auto) - fix a assume *: "a \ Field r" - hence "bij_betw f (under r a) (under r' (f a))" - using EMB unfolding embed_def by simp - moreover - {have "under r a \ Field r" - by (auto simp add: under_Field) - hence "\ b. b \ under r a \ f b = g b" - using EQ by blast - } - moreover have "f a = g a" using * EQ by auto - ultimately show "bij_betw g (under r a) (under r' (g a))" - using bij_betw_cong[of "under r a" f g "under r' (f a)"] by auto -qed + assumes "\ a. a \ Field r \ f a = g a" and "embed r r' f" + shows "embed r r' g" + by (smt (verit, del_insts) assms bij_betw_cong embed_def in_mono under_Field) lemma embed_cong[fundef_cong]: -assumes "\ a. a \ Field r \ f a = g a" -shows "embed r r' f = embed r r' g" -using assms embed_halfcong[of r f g r'] - embed_halfcong[of r g f r'] by auto + assumes "\ a. a \ Field r \ f a = g a" + shows "embed r r' f = embed r r' g" + by (metis assms embed_halfcong) lemma embedS_cong[fundef_cong]: -assumes "\ a. a \ Field r \ f a = g a" -shows "embedS r r' f = embedS r r' g" -unfolding embedS_def using assms -embed_cong[of r f g r'] bij_betw_cong[of "Field r" f g "Field r'"] by blast + assumes "\ a. a \ Field r \ f a = g a" + shows "embedS r r' f = embedS r r' g" + unfolding embedS_def using assms + embed_cong[of r f g r'] bij_betw_cong[of "Field r" f g "Field r'"] by blast lemma iso_cong[fundef_cong]: -assumes "\ a. a \ Field r \ f a = g a" -shows "iso r r' f = iso r r' g" -unfolding iso_def using assms -embed_cong[of r f g r'] bij_betw_cong[of "Field r" f g "Field r'"] by blast + assumes "\ a. a \ Field r \ f a = g a" + shows "iso r r' f = iso r r' g" + unfolding iso_def using assms + embed_cong[of r f g r'] bij_betw_cong[of "Field r" f g "Field r'"] by blast lemma id_compat: "compat r r id" -by(auto simp add: id_def compat_def) + by(auto simp add: id_def compat_def) lemma comp_compat: -"\compat r r' f; compat r' r'' f'\ \ compat r r'' (f' o f)" -by(auto simp add: comp_def compat_def) + "\compat r r' f; compat r' r'' f'\ \ compat r r'' (f' o f)" + by(auto simp add: comp_def compat_def) corollary one_set_greater: -"(\f::'a \ 'a'. f ` A \ A' \ inj_on f A) \ (\g::'a' \ 'a. g ` A' \ A \ inj_on g A')" + "(\f::'a \ 'a'. f ` A \ A' \ inj_on f A) \ (\g::'a' \ 'a. g ` A' \ A \ inj_on g A')" proof- obtain r where "well_order_on A r" by (fastforce simp add: well_order_on) hence 1: "A = Field r \ Well_order r" - using well_order_on_Well_order by auto + using well_order_on_Well_order by auto obtain r' where 2: "well_order_on A' r'" by (fastforce simp add: well_order_on) hence 2: "A' = Field r' \ Well_order r'" - using well_order_on_Well_order by auto + using well_order_on_Well_order by auto hence "(\f. embed r r' f) \ (\g. embed r' r g)" - using 1 2 by (auto simp add: wellorders_totally_ordered) + using 1 2 by (auto simp add: wellorders_totally_ordered) moreover {fix f assume "embed r r' f" - hence "f`A \ A' \ inj_on f A" - using 1 2 by (auto simp add: embed_Field embed_inj_on) + hence "f`A \ A' \ inj_on f A" + using 1 2 by (auto simp add: embed_Field embed_inj_on) } moreover {fix g assume "embed r' r g" - hence "g`A' \ A \ inj_on g A'" - using 1 2 by (auto simp add: embed_Field embed_inj_on) + hence "g`A' \ A \ inj_on g A'" + using 1 2 by (auto simp add: embed_Field embed_inj_on) } ultimately show ?thesis by blast qed corollary one_type_greater: -"(\f::'a \ 'a'. inj f) \ (\g::'a' \ 'a. inj g)" -using one_set_greater[of UNIV UNIV] by auto + "(\f::'a \ 'a'. inj f) \ (\g::'a' \ 'a. inj g)" + using one_set_greater[of UNIV UNIV] by auto subsection \Uniqueness of embeddings\ lemma comp_embedS: -assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" - and EMB: "embedS r r' f" and EMB': "embedS r' r'' f'" -shows "embedS r r'' (f' o f)" -proof- - have "embed r' r'' f'" using EMB' unfolding embedS_def by simp - thus ?thesis using assms by (auto simp add: embedS_comp_embed) -qed + assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''" + and EMB: "embedS r r' f" and EMB': "embedS r' r'' f'" + shows "embedS r r'' (f' o f)" + using EMB EMB' WELL WELL' embedS_comp_embed embedS_def by blast lemma iso_iff4: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "iso r r' f = (embed r r' f \ embed r' r (inv_into (Field r) f))" -using assms embed_bothWays_iso -by(unfold iso_def, auto simp add: inv_into_Field_embed_bij_betw) + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "iso r r' f = (embed r r' f \ embed r' r (inv_into (Field r) f))" + using assms embed_bothWays_iso + by(unfold iso_def, auto simp add: inv_into_Field_embed_bij_betw) lemma embed_embedS_iso: -"embed r r' f = (embedS r r' f \ iso r r' f)" -unfolding embedS_def iso_def by blast + "embed r r' f = (embedS r r' f \ iso r r' f)" + unfolding embedS_def iso_def by blast lemma not_embedS_iso: -"\ (embedS r r' f \ iso r r' f)" -unfolding embedS_def iso_def by blast + "\ (embedS r r' f \ iso r r' f)" + unfolding embedS_def iso_def by blast lemma embed_embedS_iff_not_iso: -assumes "embed r r' f" -shows "embedS r r' f = (\ iso r r' f)" -using assms unfolding embedS_def iso_def by blast + assumes "embed r r' f" + shows "embedS r r' f = (\ iso r r' f)" + using assms unfolding embedS_def iso_def by blast lemma iso_inv_into: -assumes WELL: "Well_order r" and ISO: "iso r r' f" -shows "iso r' r (inv_into (Field r) f)" -using assms unfolding iso_def -using bij_betw_inv_into inv_into_Field_embed_bij_betw by blast + assumes WELL: "Well_order r" and ISO: "iso r r' f" + shows "iso r' r (inv_into (Field r) f)" + by (meson ISO WELL bij_betw_inv_into inv_into_Field_embed_bij_betw iso_def iso_iff iso_iff2) lemma embedS_or_iso: -assumes WELL: "Well_order r" and WELL': "Well_order r'" -shows "(\g. embedS r r' g) \ (\h. embedS r' r h) \ (\f. iso r r' f)" -proof- - {fix f assume *: "embed r r' f" - {assume "bij_betw f (Field r) (Field r')" - hence ?thesis using * by (auto simp add: iso_def) - } - moreover - {assume "\ bij_betw f (Field r) (Field r')" - hence ?thesis using * by (auto simp add: embedS_def) - } - ultimately have ?thesis by auto - } - moreover - {fix f assume *: "embed r' r f" - {assume "bij_betw f (Field r') (Field r)" - hence "iso r' r f" using * by (auto simp add: iso_def) - hence "iso r r' (inv_into (Field r') f)" - using WELL' by (auto simp add: iso_inv_into) - hence ?thesis by blast - } - moreover - {assume "\ bij_betw f (Field r') (Field r)" - hence ?thesis using * by (auto simp add: embedS_def) - } - ultimately have ?thesis by auto - } - ultimately show ?thesis using WELL WELL' - wellorders_totally_ordered[of r r'] by blast -qed + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "(\g. embedS r r' g) \ (\h. embedS r' r h) \ (\f. iso r r' f)" + by (metis WELL WELL' embed_embedS_iso embed_embedS_iso iso_iff4 wellorders_totally_ordered) end diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Cardinals/Wellorder_Relation.thy --- a/src/HOL/Cardinals/Wellorder_Relation.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Cardinals/Wellorder_Relation.thy Sun Jan 15 20:00:44 2023 +0100 @@ -8,7 +8,7 @@ section \Well-Order Relations\ theory Wellorder_Relation -imports Wellfounded_More + imports Wellfounded_More begin context wo_rel @@ -17,29 +17,29 @@ subsection \Auxiliaries\ lemma PREORD: "Preorder r" -using WELL order_on_defs[of _ r] by auto + using WELL order_on_defs[of _ r] by auto lemma PARORD: "Partial_order r" -using WELL order_on_defs[of _ r] by auto + using WELL order_on_defs[of _ r] by auto lemma cases_Total2: -"\ phi a b. \{a,b} \ Field r; ((a,b) \ r - Id \ phi a b); + "\ phi a b. \{a,b} \ Field r; ((a,b) \ r - Id \ phi a b); ((b,a) \ r - Id \ phi a b); (a = b \ phi a b)\ \ phi a b" -using TOTALS by auto + using TOTALS by auto subsection \Well-founded induction and recursion adapted to non-strict well-order relations\ lemma worec_unique_fixpoint: -assumes ADM: "adm_wo H" and fp: "f = H f" -shows "f = worec H" + assumes ADM: "adm_wo H" and fp: "f = H f" + shows "f = worec H" proof- have "adm_wf (r - Id) H" - unfolding adm_wf_def - using ADM adm_wo_def[of H] underS_def[of r] by auto + unfolding adm_wf_def + using ADM adm_wo_def[of H] underS_def[of r] by auto hence "f = wfrec (r - Id) H" - using fp WF wfrec_unique_fixpoint[of "r - Id" H] by simp + using fp WF wfrec_unique_fixpoint[of "r - Id" H] by simp thus ?thesis unfolding worec_def . qed @@ -47,121 +47,66 @@ subsubsection \Properties of max2\ lemma max2_iff: -assumes "a \ Field r" and "b \ Field r" -shows "((max2 a b, c) \ r) = ((a,c) \ r \ (b,c) \ r)" + assumes "a \ Field r" and "b \ Field r" + shows "((max2 a b, c) \ r) = ((a,c) \ r \ (b,c) \ r)" proof assume "(max2 a b, c) \ r" thus "(a,c) \ r \ (b,c) \ r" - using assms max2_greater[of a b] TRANS trans_def[of r] by blast + using assms max2_greater[of a b] TRANS trans_def[of r] by blast next assume "(a,c) \ r \ (b,c) \ r" thus "(max2 a b, c) \ r" - using assms max2_among[of a b] by auto + using assms max2_among[of a b] by auto qed subsubsection \Properties of minim\ lemma minim_Under: -"\B \ Field r; B \ {}\ \ minim B \ Under B" -by(auto simp add: Under_def minim_inField minim_least) + "\B \ Field r; B \ {}\ \ minim B \ Under B" + by(auto simp add: Under_def minim_inField minim_least) lemma equals_minim_Under: -"\B \ Field r; a \ B; a \ Under B\ + "\B \ Field r; a \ B; a \ Under B\ \ a = minim B" -by(auto simp add: Under_def equals_minim) + by(auto simp add: Under_def equals_minim) lemma minim_iff_In_Under: -assumes SUB: "B \ Field r" and NE: "B \ {}" -shows "(a = minim B) = (a \ B \ a \ Under B)" + assumes SUB: "B \ Field r" and NE: "B \ {}" + shows "(a = minim B) = (a \ B \ a \ Under B)" proof assume "a = minim B" thus "a \ B \ a \ Under B" - using assms minim_in minim_Under by simp + using assms minim_in minim_Under by simp next assume "a \ B \ a \ Under B" thus "a = minim B" - using assms equals_minim_Under by simp + using assms equals_minim_Under by simp qed lemma minim_Under_under: -assumes NE: "A \ {}" and SUB: "A \ Field r" -shows "Under A = under (minim A)" + assumes NE: "A \ {}" and SUB: "A \ Field r" + shows "Under A = under (minim A)" proof- - (* Preliminary facts *) - have 1: "minim A \ A" - using assms minim_in by auto - have 2: "\x \ A. (minim A, x) \ r" - using assms minim_least by auto - (* Main proof *) - have "Under A \ under (minim A)" - proof - fix x assume "x \ Under A" - with 1 Under_def[of r] have "(x,minim A) \ r" by auto - thus "x \ under(minim A)" unfolding under_def by simp - qed - (* *) - moreover - (* *) - have "under (minim A) \ Under A" - proof - fix x assume "x \ under(minim A)" - hence 11: "(x,minim A) \ r" unfolding under_def by simp - hence "x \ Field r" unfolding Field_def by auto - moreover - {fix a assume "a \ A" - with 2 have "(minim A, a) \ r" by simp - with 11 have "(x,a) \ r" - using TRANS trans_def[of r] by blast - } - ultimately show "x \ Under A" by (unfold Under_def, auto) - qed - (* *) + have "minim A \ A" + using assms minim_in by auto + then have "Under A \ under (minim A)" + by (simp add: Under_decr under_Under_singl) + moreover have "under (minim A) \ Under A" + by (meson NE SUB TRANS minim_Under subset_eq under_Under_trans) ultimately show ?thesis by blast qed lemma minim_UnderS_underS: -assumes NE: "A \ {}" and SUB: "A \ Field r" -shows "UnderS A = underS (minim A)" + assumes NE: "A \ {}" and SUB: "A \ Field r" + shows "UnderS A = underS (minim A)" proof- - (* Preliminary facts *) - have 1: "minim A \ A" - using assms minim_in by auto - have 2: "\x \ A. (minim A, x) \ r" - using assms minim_least by auto - (* Main proof *) - have "UnderS A \ underS (minim A)" - proof - fix x assume "x \ UnderS A" - with 1 UnderS_def[of r] have "x \ minim A \ (x,minim A) \ r" by auto - thus "x \ underS(minim A)" unfolding underS_def by simp - qed - (* *) - moreover - (* *) - have "underS (minim A) \ UnderS A" - proof - fix x assume "x \ underS(minim A)" - hence 11: "x \ minim A \ (x,minim A) \ r" unfolding underS_def by simp - hence "x \ Field r" unfolding Field_def by auto - moreover - {fix a assume "a \ A" - with 2 have 3: "(minim A, a) \ r" by simp - with 11 have "(x,a) \ r" - using TRANS trans_def[of r] by blast - moreover - have "x \ a" - proof - assume "x = a" - with 11 3 ANTISYM antisym_def[of r] - show False by auto - qed - ultimately - have "x \ a \ (x,a) \ r" by simp - } - ultimately show "x \ UnderS A" by (unfold UnderS_def, auto) - qed - (* *) + have "minim A \ A" + using assms minim_in by auto + then have "UnderS A \ underS (minim A)" + by (simp add: UnderS_decr underS_UnderS_singl) + moreover have "underS (minim A) \ UnderS A" + by (meson ANTISYM NE SUB TRANS minim_Under subset_eq underS_Under_trans) ultimately show ?thesis by blast qed @@ -169,150 +114,103 @@ subsubsection \Properties of supr\ lemma supr_Above: -assumes SUB: "B \ Field r" and ABOVE: "Above B \ {}" -shows "supr B \ Above B" -proof(unfold supr_def) - have "Above B \ Field r" - using Above_Field[of r] by auto - thus "minim (Above B) \ Above B" - using assms by (simp add: minim_in) -qed + assumes "Above B \ {}" + shows "supr B \ Above B" + by (simp add: assms Above_Field minim_in supr_def) lemma supr_greater: -assumes SUB: "B \ Field r" and ABOVE: "Above B \ {}" and - IN: "b \ B" -shows "(b, supr B) \ r" -proof- - from assms supr_Above - have "supr B \ Above B" by simp - with IN Above_def[of r] show ?thesis by simp -qed + assumes "Above B \ {}" "b \ B" + shows "(b, supr B) \ r" + using assms Above_def supr_Above by fastforce lemma supr_least_Above: -assumes SUB: "B \ Field r" and - ABOVE: "a \ Above B" -shows "(supr B, a) \ r" -proof(unfold supr_def) - have "Above B \ Field r" - using Above_Field[of r] by auto - thus "(minim (Above B), a) \ r" - using assms minim_least - by simp -qed + assumes "a \ Above B" + shows "(supr B, a) \ r" + by (simp add: assms Above_Field minim_least supr_def) lemma supr_least: -"\B \ Field r; a \ Field r; (\ b. b \ B \ (b,a) \ r)\ + "\B \ Field r; a \ Field r; (\ b. b \ B \ (b,a) \ r)\ \ (supr B, a) \ r" -by(auto simp add: supr_least_Above Above_def) + by(auto simp add: supr_least_Above Above_def) lemma equals_supr_Above: -assumes SUB: "B \ Field r" and ABV: "a \ Above B" and - MINIM: "\ a'. a' \ Above B \ (a,a') \ r" -shows "a = supr B" -proof(unfold supr_def) - have "Above B \ Field r" - using Above_Field[of r] by auto - thus "a = minim (Above B)" - using assms equals_minim by simp -qed + assumes "a \ Above B" "\ a'. a' \ Above B \ (a,a') \ r" + shows "a = supr B" + by (simp add: assms Above_Field equals_minim supr_def) lemma equals_supr: -assumes SUB: "B \ Field r" and IN: "a \ Field r" and - ABV: "\ b. b \ B \ (b,a) \ r" and - MINIM: "\ a'. \ a' \ Field r; \ b. b \ B \ (b,a') \ r\ \ (a,a') \ r" -shows "a = supr B" + assumes SUB: "B \ Field r" and IN: "a \ Field r" and + ABV: "\ b. b \ B \ (b,a) \ r" and + MINIM: "\ a'. \ a' \ Field r; \ b. b \ B \ (b,a') \ r\ \ (a,a') \ r" + shows "a = supr B" proof- have "a \ Above B" - unfolding Above_def using ABV IN by simp + unfolding Above_def using ABV IN by simp moreover have "\ a'. a' \ Above B \ (a,a') \ r" - unfolding Above_def using MINIM by simp + unfolding Above_def using MINIM by simp ultimately show ?thesis - using equals_supr_Above SUB by auto + using equals_supr_Above SUB by auto qed lemma supr_inField: -assumes "B \ Field r" and "Above B \ {}" -shows "supr B \ Field r" -proof- - have "supr B \ Above B" using supr_Above assms by simp - thus ?thesis using assms Above_Field[of r] by auto -qed + assumes "Above B \ {}" + shows "supr B \ Field r" + by (simp add: Above_Field assms minim_inField supr_def) lemma supr_above_Above: -assumes SUB: "B \ Field r" and ABOVE: "Above B \ {}" -shows "Above B = above (supr B)" -proof(unfold Above_def above_def, auto) - fix a assume "a \ Field r" "\b \ B. (b,a) \ r" - with supr_least assms - show "(supr B, a) \ r" by auto -next - fix b assume "(supr B, b) \ r" - thus "b \ Field r" - using REFL refl_on_def[of _ r] by auto -next - fix a b - assume 1: "(supr B, b) \ r" and 2: "a \ B" - with assms supr_greater - have "(a,supr B) \ r" by auto - thus "(a,b) \ r" - using 1 TRANS trans_def[of r] by blast -qed + assumes SUB: "B \ Field r" and ABOVE: "Above B \ {}" + shows "Above B = above (supr B)" + apply (clarsimp simp: Above_def above_def set_eq_iff) + by (meson ABOVE FieldI2 SUB TRANS supr_greater supr_least transD) lemma supr_under: -assumes IN: "a \ Field r" -shows "a = supr (under a)" + assumes "a \ Field r" + shows "a = supr (under a)" proof- have "under a \ Field r" - using under_Field[of r] by auto + using under_Field[of r] by auto moreover have "under a \ {}" - using IN Refl_under_in[of r] REFL by auto + using assms Refl_under_in[of r] REFL by auto moreover have "a \ Above (under a)" - using in_Above_under[of _ r] IN by auto + using in_Above_under[of _ r] assms by auto moreover have "\a' \ Above (under a). (a,a') \ r" - proof(unfold Above_def under_def, auto) - fix a' - assume "\aa. (aa, a) \ r \ (aa, a') \ r" - hence "(a,a) \ r \ (a,a') \ r" by blast - moreover have "(a,a) \ r" - using REFL IN by (auto simp add: refl_on_def) - ultimately - show "(a, a') \ r" by (rule mp) - qed + by (auto simp: Above_def above_def REFL Refl_under_in assms) ultimately show ?thesis - using equals_supr_Above by auto + using equals_supr_Above by auto qed subsubsection \Properties of successor\ lemma suc_least: -"\B \ Field r; a \ Field r; (\ b. b \ B \ a \ b \ (b,a) \ r)\ + "\B \ Field r; a \ Field r; (\ b. b \ B \ a \ b \ (b,a) \ r)\ \ (suc B, a) \ r" -by(auto simp add: suc_least_AboveS AboveS_def) + by(auto simp add: suc_least_AboveS AboveS_def) lemma equals_suc: -assumes SUB: "B \ Field r" and IN: "a \ Field r" and - ABVS: "\ b. b \ B \ a \ b \ (b,a) \ r" and - MINIM: "\ a'. \a' \ Field r; \ b. b \ B \ a' \ b \ (b,a') \ r\ \ (a,a') \ r" -shows "a = suc B" + assumes SUB: "B \ Field r" and IN: "a \ Field r" and + ABVS: "\ b. b \ B \ a \ b \ (b,a) \ r" and + MINIM: "\ a'. \a' \ Field r; \ b. b \ B \ a' \ b \ (b,a') \ r\ \ (a,a') \ r" + shows "a = suc B" proof- have "a \ AboveS B" - unfolding AboveS_def using ABVS IN by simp + unfolding AboveS_def using ABVS IN by simp moreover have "\ a'. a' \ AboveS B \ (a,a') \ r" - unfolding AboveS_def using MINIM by simp + unfolding AboveS_def using MINIM by simp ultimately show ?thesis - using equals_suc_AboveS SUB by auto + using equals_suc_AboveS SUB by auto qed lemma suc_above_AboveS: -assumes SUB: "B \ Field r" and - ABOVE: "AboveS B \ {}" -shows "AboveS B = above (suc B)" + assumes SUB: "B \ Field r" and + ABOVE: "AboveS B \ {}" + shows "AboveS B = above (suc B)" + using assms proof(unfold AboveS_def above_def, auto) fix a assume "a \ Field r" "\b \ B. a \ b \ (b,a) \ r" with suc_least assms @@ -320,82 +218,57 @@ next fix b assume "(suc B, b) \ r" thus "b \ Field r" - using REFL refl_on_def[of _ r] by auto + using REFL refl_on_def[of _ r] by auto next fix a b assume 1: "(suc B, b) \ r" and 2: "a \ B" with assms suc_greater[of B a] have "(a,suc B) \ r" by auto thus "(a,b) \ r" - using 1 TRANS trans_def[of r] by blast + using 1 TRANS trans_def[of r] by blast next fix a - assume 1: "(suc B, a) \ r" and 2: "a \ B" - with assms suc_greater[of B a] - have "(a,suc B) \ r" by auto - moreover have "suc B \ Field r" - using assms suc_inField by simp - ultimately have "a = suc B" - using 1 2 SUB ANTISYM antisym_def[of r] by auto + assume "(suc B, a) \ r" and 2: "a \ B" thus False - using assms suc_greater[of B a] 2 by auto + by (metis ABOVE ANTISYM SUB antisymD suc_greater) qed lemma suc_singl_pred: -assumes IN: "a \ Field r" and ABOVE_NE: "aboveS a \ {}" and - REL: "(a',suc {a}) \ r" and DIFF: "a' \ suc {a}" -shows "a' = a \ (a',a) \ r" + assumes IN: "a \ Field r" and ABOVE_NE: "aboveS a \ {}" and + REL: "(a',suc {a}) \ r" and DIFF: "a' \ suc {a}" + shows "a' = a \ (a',a) \ r" proof- have *: "suc {a} \ Field r \ a' \ Field r" - using WELL REL well_order_on_domain by metis + using WELL REL well_order_on_domain by metis {assume **: "a' \ a" - hence "(a,a') \ r \ (a',a) \ r" - using TOTAL IN * by (auto simp add: total_on_def) - moreover - {assume "(a,a') \ r" - with ** * assms WELL suc_least[of "{a}" a'] - have "(suc {a},a') \ r" by auto - with REL DIFF * ANTISYM antisym_def[of r] - have False by simp - } - ultimately have "(a',a) \ r" - by blast + hence "(a,a') \ r \ (a',a) \ r" + using TOTAL IN * by (auto simp add: total_on_def) + moreover + {assume "(a,a') \ r" + with ** * assms WELL suc_least[of "{a}" a'] + have "(suc {a},a') \ r" by auto + with REL DIFF * ANTISYM antisym_def[of r] + have False by simp + } + ultimately have "(a',a) \ r" + by blast } thus ?thesis by blast qed lemma under_underS_suc: -assumes IN: "a \ Field r" and ABV: "aboveS a \ {}" -shows "underS (suc {a}) = under a" -proof- - have 1: "AboveS {a} \ {}" - using ABV aboveS_AboveS_singl[of r] by auto - have 2: "a \ suc {a} \ (a,suc {a}) \ r" - using suc_greater[of "{a}" a] IN 1 by auto - (* *) + assumes IN: "a \ Field r" and ABV: "aboveS a \ {}" + shows "underS (suc {a}) = under a" +proof - + have "AboveS {a} \ {}" + using ABV aboveS_AboveS_singl[of r] by auto + then have 2: "a \ suc {a} \ (a,suc {a}) \ r" + using suc_greater[of "{a}" a] IN by auto have "underS (suc {a}) \ under a" - proof(unfold underS_def under_def, auto) - fix x assume *: "x \ suc {a}" and **: "(x,suc {a}) \ r" - with suc_singl_pred[of a x] IN ABV - have "x = a \ (x,a) \ r" by auto - with REFL refl_on_def[of _ r] IN - show "(x,a) \ r" by auto - qed - (* *) + using ABV IN REFL Refl_under_in underS_E under_def wo_rel.suc_singl_pred wo_rel_axioms by fastforce moreover - (* *) have "under a \ underS (suc {a})" - proof(unfold underS_def under_def, auto) - assume "(suc {a}, a) \ r" - with 2 ANTISYM antisym_def[of r] - show False by auto - next - fix x assume *: "(x,a) \ r" - with 2 TRANS trans_def[of r] - show "(x,suc {a}) \ r" by blast - (* blast is often better than auto/auto for transitivity-like properties *) - qed - (* *) + by (simp add: "2" ANTISYM TRANS subsetI underS_I under_underS_trans) ultimately show ?thesis by blast qed @@ -403,113 +276,49 @@ subsubsection \Properties of order filters\ lemma ofilter_Under[simp]: -assumes "A \ Field r" -shows "ofilter(Under A)" -proof(unfold ofilter_def, auto) - fix x assume "x \ Under A" - thus "x \ Field r" - using Under_Field[of r] assms by auto -next - fix a x - assume "a \ Under A" and "x \ under a" - thus "x \ Under A" - using TRANS under_Under_trans[of r] by auto -qed + assumes "A \ Field r" + shows "ofilter(Under A)" + by (clarsimp simp: ofilter_def) (meson TRANS Under_Field subset_eq under_Under_trans) lemma ofilter_UnderS[simp]: -assumes "A \ Field r" -shows "ofilter(UnderS A)" -proof(unfold ofilter_def, auto) - fix x assume "x \ UnderS A" - thus "x \ Field r" - using UnderS_Field[of r] assms by auto -next - fix a x - assume "a \ UnderS A" and "x \ under a" - thus "x \ UnderS A" - using TRANS ANTISYM under_UnderS_trans[of r] by auto -qed + assumes "A \ Field r" + shows "ofilter(UnderS A)" + by (clarsimp simp: ofilter_def) (meson ANTISYM TRANS UnderS_Field subset_eq under_UnderS_trans) lemma ofilter_Int[simp]: "\ofilter A; ofilter B\ \ ofilter(A Int B)" -unfolding ofilter_def by blast + unfolding ofilter_def by blast lemma ofilter_Un[simp]: "\ofilter A; ofilter B\ \ ofilter(A \ B)" -unfolding ofilter_def by blast + unfolding ofilter_def by blast lemma ofilter_INTER: -"\I \ {}; \ i. i \ I \ ofilter(A i)\ \ ofilter (\i \ I. A i)" -unfolding ofilter_def by blast + "\I \ {}; \ i. i \ I \ ofilter(A i)\ \ ofilter (\i \ I. A i)" + unfolding ofilter_def by blast lemma ofilter_Inter: -"\S \ {}; \ A. A \ S \ ofilter A\ \ ofilter (\S)" -unfolding ofilter_def by blast + "\S \ {}; \ A. A \ S \ ofilter A\ \ ofilter (\S)" + unfolding ofilter_def by blast lemma ofilter_Union: -"(\ A. A \ S \ ofilter A) \ ofilter (\S)" -unfolding ofilter_def by blast + "(\ A. A \ S \ ofilter A) \ ofilter (\S)" + unfolding ofilter_def by blast lemma ofilter_under_Union: -"ofilter A \ A = \{under a| a. a \ A}" -using ofilter_under_UNION [of A] by auto + "ofilter A \ A = \{under a| a. a \ A}" + using ofilter_under_UNION [of A] by auto subsubsection \Other properties\ lemma Trans_Under_regressive: -assumes NE: "A \ {}" and SUB: "A \ Field r" -shows "Under(Under A) \ Under A" -proof - let ?a = "minim A" - (* Preliminary facts *) - have 1: "minim A \ Under A" - using assms minim_Under by auto - have 2: "\y \ A. (minim A, y) \ r" - using assms minim_least by auto - (* Main proof *) - fix x assume "x \ Under(Under A)" - with 1 have 1: "(x,minim A) \ r" - using Under_def[of r] by auto - with Field_def have "x \ Field r" by fastforce - moreover - {fix y assume *: "y \ A" - hence "(x,y) \ r" - using 1 2 TRANS trans_def[of r] by blast - with Field_def have "(x,y) \ r" by auto - } - ultimately - show "x \ Under A" unfolding Under_def by auto -qed + assumes NE: "A \ {}" and SUB: "A \ Field r" + shows "Under(Under A) \ Under A" + by (metis INT_E NE REFL Refl_under_Under SUB empty_iff minim_Under minim_Under_under subsetI) lemma ofilter_suc_Field: -assumes OF: "ofilter A" and NE: "A \ Field r" -shows "ofilter (A \ {suc A})" -proof- - (* Preliminary facts *) - have 1: "A \ Field r" using OF ofilter_def by auto - hence 2: "AboveS A \ {}" - using ofilter_AboveS_Field NE OF by blast - from 1 2 suc_inField - have 3: "suc A \ Field r" by auto - (* Main proof *) - show ?thesis - proof(unfold ofilter_def, auto simp add: 1 3) - fix a x - assume "a \ A" "x \ under a" "x \ A" - with OF ofilter_def have False by auto - thus "x = suc A" by simp - next - fix x assume *: "x \ under (suc A)" and **: "x \ A" - hence "x \ Field r" using under_def Field_def by fastforce - with ** have "x \ AboveS A" - using ofilter_AboveS_Field[of A] OF by auto - hence "(suc A,x) \ r" - using suc_least_AboveS by auto - moreover - have "(x,suc A) \ r" using * under_def[of r] by auto - ultimately show "x = suc A" - using ANTISYM antisym_def[of r] by auto - qed -qed + assumes OF: "ofilter A" and NE: "A \ Field r" + shows "ofilter (A \ {suc A})" + by (metis NE OF REFL Refl_under_underS ofilter_underS_Field suc_underS under_ofilter) (* FIXME: needed? *) declare diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Relation.thy --- a/src/HOL/Relation.thy Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Relation.thy Sun Jan 15 20:00:44 2023 +0100 @@ -1247,6 +1247,9 @@ definition Field :: "'a rel \ 'a set" where "Field r = Domain r \ Range r" +lemma Field_iff: "x \ Field r \ (\y. (x,y) \ r \ (y,x) \ r)" + by (auto simp: Field_def) + lemma FieldI1: "(i, j) \ R \ i \ Field R" unfolding Field_def by blast diff -r f327ae3cab2a -r d3de24c50b08 src/HOL/Tools/Mirabelle/mirabelle_sledgehammer.ML --- a/src/HOL/Tools/Mirabelle/mirabelle_sledgehammer.ML Sun Jan 15 20:00:37 2023 +0100 +++ b/src/HOL/Tools/Mirabelle/mirabelle_sledgehammer.ML Sun Jan 15 20:00:44 2023 +0100 @@ -305,12 +305,14 @@ exhaustive_preplay proof_method_from_msg thy_index trivial pos st = let val thy = Proof.theory_of st - val thy_name = Context.theory_name thy + val thy_long_name = Context.theory_long_name thy + val session_name = Long_Name.qualifier thy_long_name + val thy_name = Long_Name.base_name thy_long_name val triv_str = if trivial then "[T] " else "" val keep = if keep_probs orelse keep_proofs then let val subdir = StringCvt.padLeft #"0" 4 (string_of_int thy_index) ^ "_" ^ thy_name in - Path.append output_dir (Path.basic subdir) + Path.append (Path.append output_dir (Path.basic session_name)) (Path.basic subdir) |> Isabelle_System.make_directory |> Path.implode |> (fn dir => SOME (dir, keep_probs, keep_proofs))