diff --git a/src/HOL/Algebra/Free_Abelian_Groups.thy b/src/HOL/Algebra/Free_Abelian_Groups.thy --- a/src/HOL/Algebra/Free_Abelian_Groups.thy +++ b/src/HOL/Algebra/Free_Abelian_Groups.thy @@ -1,755 +1,755 @@ section\Free Abelian Groups\ theory Free_Abelian_Groups imports Product_Groups FiniteProduct "HOL-Cardinals.Cardinal_Arithmetic" "HOL-Library.Countable_Set" "HOL-Library.Poly_Mapping" "HOL-Library.Equipollence" begin (*Move? But where?*) lemma eqpoll_Fpow: assumes "infinite A" shows "Fpow A \ A" unfolding eqpoll_iff_card_of_ordIso by (metis assms card_of_Fpow_infinite) lemma infinite_iff_card_of_countable: "\countable B; infinite B\ \ infinite A \ ( |B| \o |A| )" unfolding infinite_iff_countable_subset card_of_ordLeq countable_def by (force intro: card_of_ordLeqI ordLeq_transitive) lemma iso_imp_eqpoll_carrier: "G \ H \ carrier G \ carrier H" by (auto simp: is_iso_def iso_def eqpoll_def) subsection\Generalised finite product\ definition gfinprod :: "[('b, 'm) monoid_scheme, 'a \ 'b, 'a set] \ 'b" where "gfinprod G f A = (if finite {x \ A. f x \ \\<^bsub>G\<^esub>} then finprod G f {x \ A. f x \ \\<^bsub>G\<^esub>} else \\<^bsub>G\<^esub>)" context comm_monoid begin lemma gfinprod_closed [simp]: "f \ A \ carrier G \ gfinprod G f A \ carrier G" unfolding gfinprod_def by (auto simp: image_subset_iff_funcset intro: finprod_closed) lemma gfinprod_cong: "\A = B; f \ B \ carrier G; \i. i \ B =simp=> f i = g i\ \ gfinprod G f A = gfinprod G g B" unfolding gfinprod_def by (auto simp: simp_implies_def cong: conj_cong intro: finprod_cong) lemma gfinprod_eq_finprod [simp]: "\finite A; f \ A \ carrier G\ \ gfinprod G f A = finprod G f A" by (auto simp: gfinprod_def intro: finprod_mono_neutral_cong_left) lemma gfinprod_insert [simp]: assumes "finite {x \ A. f x \ \\<^bsub>G\<^esub>}" "f \ A \ carrier G" "f i \ carrier G" shows "gfinprod G f (insert i A) = (if i \ A then gfinprod G f A else f i \ gfinprod G f A)" proof - have f: "f \ {x \ A. f x \ \} \ carrier G" using assms by (auto simp: image_subset_iff_funcset) have "{x. x = i \ f x \ \ \ x \ A \ f x \ \} = (if f i = \ then {x \ A. f x \ \} else insert i {x \ A. f x \ \})" by auto then show ?thesis using assms unfolding gfinprod_def by (simp add: conj_disj_distribR insert_absorb f split: if_split_asm) qed lemma gfinprod_distrib: assumes fin: "finite {x \ A. f x \ \\<^bsub>G\<^esub>}" "finite {x \ A. g x \ \\<^bsub>G\<^esub>}" and "f \ A \ carrier G" "g \ A \ carrier G" shows "gfinprod G (\i. f i \ g i) A = gfinprod G f A \ gfinprod G g A" proof - have "finite {x \ A. f x \ g x \ \}" by (auto intro: finite_subset [OF _ finite_UnI [OF fin]]) then have "gfinprod G (\i. f i \ g i) A = gfinprod G (\i. f i \ g i) ({i \ A. f i \ \\<^bsub>G\<^esub>} \ {i \ A. g i \ \\<^bsub>G\<^esub>})" unfolding gfinprod_def using assms by (force intro: finprod_mono_neutral_cong) also have "\ = gfinprod G f A \ gfinprod G g A" proof - have "finprod G f ({i \ A. f i \ \\<^bsub>G\<^esub>} \ {i \ A. g i \ \\<^bsub>G\<^esub>}) = gfinprod G f A" "finprod G g ({i \ A. f i \ \\<^bsub>G\<^esub>} \ {i \ A. g i \ \\<^bsub>G\<^esub>}) = gfinprod G g A" using assms by (auto simp: gfinprod_def intro: finprod_mono_neutral_cong_right) moreover have "(\i. f i \ g i) \ {i \ A. f i \ \} \ {i \ A. g i \ \} \ carrier G" using assms by (force simp: image_subset_iff_funcset) ultimately show ?thesis using assms apply simp apply (subst finprod_multf, auto) done qed finally show ?thesis . qed lemma gfinprod_mono_neutral_cong_left: assumes "A \ B" and 1: "\i. i \ B - A \ h i = \" and gh: "\x. x \ A \ g x = h x" and h: "h \ B \ carrier G" shows "gfinprod G g A = gfinprod G h B" proof (cases "finite {x \ B. h x \ \}") case True then have "finite {x \ A. h x \ \}" apply (rule rev_finite_subset) using \A \ B\ by auto with True assms show ?thesis apply (simp add: gfinprod_def cong: conj_cong) apply (auto intro!: finprod_mono_neutral_cong_left) done next case False have "{x \ B. h x \ \} \ {x \ A. h x \ \}" using 1 by auto with False have "infinite {x \ A. h x \ \}" using infinite_super by blast with False assms show ?thesis by (simp add: gfinprod_def cong: conj_cong) qed lemma gfinprod_mono_neutral_cong_right: assumes "A \ B" "\i. i \ B - A \ g i = \" "\x. x \ A \ g x = h x" "g \ B \ carrier G" shows "gfinprod G g B = gfinprod G h A" using assms by (auto intro!: gfinprod_mono_neutral_cong_left [symmetric]) lemma gfinprod_mono_neutral_cong: assumes [simp]: "finite B" "finite A" and *: "\i. i \ B - A \ h i = \" "\i. i \ A - B \ g i = \" and gh: "\x. x \ A \ B \ g x = h x" and g: "g \ A \ carrier G" and h: "h \ B \ carrier G" shows "gfinprod G g A = gfinprod G h B" proof- have "gfinprod G g A = gfinprod G g (A \ B)" by (rule gfinprod_mono_neutral_cong_right) (use assms in auto) also have "\ = gfinprod G h (A \ B)" by (rule gfinprod_cong) (use assms in auto) also have "\ = gfinprod G h B" by (rule gfinprod_mono_neutral_cong_left) (use assms in auto) finally show ?thesis . qed end lemma (in comm_group) hom_group_sum: assumes hom: "\i. i \ I \ f i \ hom (A i) G" and grp: "\i. i \ I \ group (A i)" shows "(\x. gfinprod G (\i. (f i) (x i)) I) \ hom (sum_group I A) G" unfolding hom_def proof (intro CollectI conjI ballI) show "(\x. gfinprod G (\i. f i (x i)) I) \ carrier (sum_group I A) \ carrier G" using assms by (force simp: hom_def carrier_sum_group intro: gfinprod_closed simp flip: image_subset_iff_funcset) next fix x y assume x: "x \ carrier (sum_group I A)" and y: "y \ carrier (sum_group I A)" then have finx: "finite {i \ I. x i \ \\<^bsub>A i\<^esub>}" and finy: "finite {i \ I. y i \ \\<^bsub>A i\<^esub>}" using assms by (auto simp: carrier_sum_group) have finfx: "finite {i \ I. f i (x i) \ \}" using assms by (auto simp: is_group hom_one [OF hom] intro: finite_subset [OF _ finx]) have finfy: "finite {i \ I. f i (y i) \ \}" using assms by (auto simp: is_group hom_one [OF hom] intro: finite_subset [OF _ finy]) have carr: "f i (x i) \ carrier G" "f i (y i) \ carrier G" if "i \ I" for i using hom_carrier [OF hom] that x y assms by (fastforce simp add: carrier_sum_group)+ have lam: "(\i. f i ( x i \\<^bsub>A i\<^esub> y i)) \ I \ carrier G" using x y assms by (auto simp: hom_def carrier_sum_group PiE_def Pi_def) have lam': "(\i. f i (if i \ I then x i \\<^bsub>A i\<^esub> y i else undefined)) \ I \ carrier G" by (simp add: lam Pi_cong) with lam x y assms show "gfinprod G (\i. f i ((x \\<^bsub>sum_group I A\<^esub> y) i)) I = gfinprod G (\i. f i (x i)) I \ gfinprod G (\i. f i (y i)) I" by (simp add: carrier_sum_group PiE_def Pi_def hom_mult [OF hom] gfinprod_distrib finfx finfy carr cong: gfinprod_cong) qed subsection\Free Abelian groups on a set, using the "frag" type constructor. \ definition free_Abelian_group :: "'a set \ ('a \\<^sub>0 int) monoid" where "free_Abelian_group S = \carrier = {c. Poly_Mapping.keys c \ S}, monoid.mult = (+), one = 0\" lemma group_free_Abelian_group [simp]: "group (free_Abelian_group S)" proof - have "\x. Poly_Mapping.keys x \ S \ x \ Units (free_Abelian_group S)" unfolding free_Abelian_group_def Units_def by clarsimp (metis eq_neg_iff_add_eq_0 neg_eq_iff_add_eq_0 keys_minus) then show ?thesis unfolding free_Abelian_group_def by unfold_locales (auto simp: dest: subsetD [OF keys_add]) qed lemma carrier_free_Abelian_group_iff [simp]: shows "x \ carrier (free_Abelian_group S) \ Poly_Mapping.keys x \ S" by (auto simp: free_Abelian_group_def) lemma one_free_Abelian_group [simp]: "\\<^bsub>free_Abelian_group S\<^esub> = 0" by (auto simp: free_Abelian_group_def) lemma mult_free_Abelian_group [simp]: "x \\<^bsub>free_Abelian_group S\<^esub> y = x + y" by (auto simp: free_Abelian_group_def) lemma inv_free_Abelian_group [simp]: "Poly_Mapping.keys x \ S \ inv\<^bsub>free_Abelian_group S\<^esub> x = -x" by (rule group.inv_equality [OF group_free_Abelian_group]) auto lemma abelian_free_Abelian_group: "comm_group(free_Abelian_group S)" apply (rule group.group_comm_groupI [OF group_free_Abelian_group]) by (simp add: free_Abelian_group_def) lemma pow_free_Abelian_group [simp]: fixes n::nat shows "Group.pow (free_Abelian_group S) x n = frag_cmul (int n) x" by (induction n) (auto simp: nat_pow_def free_Abelian_group_def frag_cmul_distrib) lemma int_pow_free_Abelian_group [simp]: fixes n::int assumes "Poly_Mapping.keys x \ S" shows "Group.pow (free_Abelian_group S) x n = frag_cmul n x" proof (induction n) case (nonneg n) then show ?case by (simp add: int_pow_int) next case (neg n) have "x [^]\<^bsub>free_Abelian_group S\<^esub> - int (Suc n) = inv\<^bsub>free_Abelian_group S\<^esub> (x [^]\<^bsub>free_Abelian_group S\<^esub> int (Suc n))" by (rule group.int_pow_neg [OF group_free_Abelian_group]) (use assms in \simp add: free_Abelian_group_def\) also have "\ = frag_cmul (- int (Suc n)) x" by (metis assms inv_free_Abelian_group pow_free_Abelian_group int_pow_int minus_frag_cmul order_trans keys_cmul) finally show ?case . qed lemma frag_of_in_free_Abelian_group [simp]: "frag_of x \ carrier(free_Abelian_group S) \ x \ S" by simp lemma free_Abelian_group_induct: assumes major: "Poly_Mapping.keys x \ S" and minor: "P(0)" "\x y. \Poly_Mapping.keys x \ S; Poly_Mapping.keys y \ S; P x; P y\ \ P(x-y)" "\a. a \ S \ P(frag_of a)" shows "P x" proof - have "Poly_Mapping.keys x \ S \ P x" using major proof (induction x rule: frag_induction) case (diff a b) then show ?case by (meson Un_least minor(2) order.trans keys_diff) qed (auto intro: minor) then show ?thesis .. qed lemma sum_closed_free_Abelian_group: "(\i. i \ I \ x i \ carrier (free_Abelian_group S)) \ sum x I \ carrier (free_Abelian_group S)" apply (induction I rule: infinite_finite_induct, auto) by (metis (no_types, hide_lams) UnE subsetCE keys_add) lemma (in comm_group) free_Abelian_group_universal: fixes f :: "'c \ 'a" assumes "f ` S \ carrier G" obtains h where "h \ hom (free_Abelian_group S) G" "\x. x \ S \ h(frag_of x) = f x" proof have fin: "Poly_Mapping.keys u \ S \ finite {x \ S. f x [^] poly_mapping.lookup u x \ \}" for u :: "'c \\<^sub>0 int" apply (rule finite_subset [OF _ finite_keys [of u]]) unfolding keys.rep_eq by force define h :: "('c \\<^sub>0 int) \ 'a" where "h \ \x. gfinprod G (\a. f a [^] poly_mapping.lookup x a) S" show "h \ hom (free_Abelian_group S) G" proof (rule homI) fix x y assume xy: "x \ carrier (free_Abelian_group S)" "y \ carrier (free_Abelian_group S)" then show "h (x \\<^bsub>free_Abelian_group S\<^esub> y) = h x \ h y" using assms unfolding h_def free_Abelian_group_def by (simp add: fin gfinprod_distrib image_subset_iff Poly_Mapping.lookup_add int_pow_mult cong: gfinprod_cong) qed (use assms in \force simp: free_Abelian_group_def h_def intro: gfinprod_closed\) show "h(frag_of x) = f x" if "x \ S" for x proof - have fin: "(\a. f x [^] (1::int)) \ {x} \ carrier G" "f x [^] (1::int) \ carrier G" using assms that by force+ show ?thesis by (cases " f x [^] (1::int) = \") (use assms that in \auto simp: h_def gfinprod_def finprod_singleton\) qed qed lemma eqpoll_free_Abelian_group_infinite: assumes "infinite A" shows "carrier(free_Abelian_group A) \ A" proof (rule lepoll_antisym) have "carrier (free_Abelian_group A) \ {f::'a\int. f ` A \ UNIV \ {x. f x \ 0} \ A \ finite {x. f x \ 0}}" unfolding lepoll_def by (rule_tac x="Poly_Mapping.lookup" in exI) (auto simp: poly_mapping_eqI lookup_not_eq_zero_eq_in_keys inj_onI) also have "\ \ Fpow (A \ (UNIV::int set))" by (rule lepoll_restricted_funspace) also have "\ \ A \ (UNIV::int set)" proof (rule eqpoll_Fpow) show "infinite (A \ (UNIV::int set))" using assms finite_cartesian_productD1 by fastforce qed also have "\ \ A" unfolding eqpoll_iff_card_of_ordIso proof - have "|A \ (UNIV::int set)| <=o |A|" by (simp add: assms card_of_Times_ordLeq_infinite flip: infinite_iff_card_of_countable) moreover have "|A| \o |A \ (UNIV::int set)|" by simp ultimately have "|A| *c |(UNIV::int set)| =o |A|" by (simp add: cprod_def ordIso_iff_ordLeq) then show "|A \ (UNIV::int set)| =o |A|" by (metis Times_cprod ordIso_transitive) qed finally show "carrier (free_Abelian_group A) \ A" . have "inj_on frag_of A" by (simp add: frag_of_eq inj_on_def) moreover have "frag_of ` A \ carrier (free_Abelian_group A)" by (simp add: image_subsetI) ultimately show "A \ carrier (free_Abelian_group A)" by (force simp: lepoll_def) qed proposition (in comm_group) eqpoll_homomorphisms_from_free_Abelian_group: "{f. f \ extensional (carrier(free_Abelian_group S)) \ f \ hom (free_Abelian_group S) G} \ (S \\<^sub>E carrier G)" (is "?lhs \ ?rhs") unfolding eqpoll_def bij_betw_def proof (intro exI conjI) let ?f = "\f. restrict (f \ frag_of) S" show "inj_on ?f ?lhs" proof (clarsimp simp: inj_on_def) fix g h assume g: "g \ extensional (carrier (free_Abelian_group S))" "g \ hom (free_Abelian_group S) G" and h: "h \ extensional (carrier (free_Abelian_group S))" "h \ hom (free_Abelian_group S) G" and eq: "restrict (g \ frag_of) S = restrict (h \ frag_of) S" have 0: "0 \ carrier (free_Abelian_group S)" by simp interpret hom_g: group_hom "free_Abelian_group S" G g using g by (auto simp: group_hom_def group_hom_axioms_def is_group) interpret hom_h: group_hom "free_Abelian_group S" G h using h by (auto simp: group_hom_def group_hom_axioms_def is_group) have "Poly_Mapping.keys c \ S \ Poly_Mapping.keys c \ S \ g c = h c" for c proof (induction c rule: frag_induction) case zero show ?case using hom_g.hom_one hom_h.hom_one by auto next case (one x) then show ?case using eq by (simp add: fun_eq_iff) (metis comp_def) next case (diff a b) then show ?case using hom_g.hom_mult hom_h.hom_mult hom_g.hom_inv hom_h.hom_inv apply (auto simp: dest: subsetD [OF keys_diff]) by (metis keys_minus uminus_add_conv_diff) qed then show "g = h" by (meson g h carrier_free_Abelian_group_iff extensionalityI) qed have "f \ (\f. restrict (f \ frag_of) S) ` {f \ extensional (carrier (free_Abelian_group S)). f \ hom (free_Abelian_group S) G}" if f: "f \ S \\<^sub>E carrier G" for f :: "'c \ 'a" proof - obtain h where h: "h \ hom (free_Abelian_group S) G" "\x. x \ S \ h(frag_of x) = f x" proof (rule free_Abelian_group_universal) show "f ` S \ carrier G" using f by blast qed auto let ?h = "restrict h (carrier (free_Abelian_group S))" show ?thesis proof show "f = restrict (?h \ frag_of) S" using f by (force simp: h) show "?h \ {f \ extensional (carrier (free_Abelian_group S)). f \ hom (free_Abelian_group S) G}" using h by (auto simp: hom_def dest!: subsetD [OF keys_add]) qed qed then show "?f ` ?lhs = S \\<^sub>E carrier G" by (auto simp: hom_def Ball_def Pi_def) qed lemma hom_frag_minus: assumes "h \ hom (free_Abelian_group S) (free_Abelian_group T)" "Poly_Mapping.keys a \ S" shows "h (-a) = - (h a)" proof - have "Poly_Mapping.keys (h a) \ T" by (meson assms carrier_free_Abelian_group_iff hom_in_carrier) then show ?thesis by (metis (no_types) assms carrier_free_Abelian_group_iff group_free_Abelian_group group_hom.hom_inv group_hom_axioms_def group_hom_def inv_free_Abelian_group) qed lemma hom_frag_add: assumes "h \ hom (free_Abelian_group S) (free_Abelian_group T)" "Poly_Mapping.keys a \ S" "Poly_Mapping.keys b \ S" shows "h (a+b) = h a + h b" proof - have "Poly_Mapping.keys (h a) \ T" by (meson assms carrier_free_Abelian_group_iff hom_in_carrier) moreover have "Poly_Mapping.keys (h b) \ T" by (meson assms carrier_free_Abelian_group_iff hom_in_carrier) ultimately show ?thesis using assms hom_mult by fastforce qed lemma hom_frag_diff: assumes "h \ hom (free_Abelian_group S) (free_Abelian_group T)" "Poly_Mapping.keys a \ S" "Poly_Mapping.keys b \ S" shows "h (a-b) = h a - h b" by (metis (no_types, lifting) assms diff_conv_add_uminus hom_frag_add hom_frag_minus keys_minus) proposition isomorphic_free_Abelian_groups: "free_Abelian_group S \ free_Abelian_group T \ S \ T" (is "(?FS \ ?FT) = ?rhs") proof interpret S: group "?FS" by simp interpret T: group "?FT" by simp interpret G2: comm_group "integer_mod_group 2" by (rule abelian_integer_mod_group) let ?Two = "{0..<2::int}" have [simp]: "\ ?Two \ {a}" for a by (simp add: subset_iff) presburger assume L: "?FS \ ?FT" let ?HS = "{h \ extensional (carrier ?FS). h \ hom ?FS (integer_mod_group 2)}" let ?HT = "{h \ extensional (carrier ?FT). h \ hom ?FT (integer_mod_group 2)}" have "S \\<^sub>E ?Two \ ?HS" apply (rule eqpoll_sym) using G2.eqpoll_homomorphisms_from_free_Abelian_group by (simp add: carrier_integer_mod_group) also have "\ \ ?HT" proof - obtain f g where "group_isomorphisms ?FS ?FT f g" using L S.iso_iff_group_isomorphisms by (force simp: is_iso_def) then have f: "f \ hom ?FS ?FT" and g: "g \ hom ?FT ?FS" and gf: "\x \ carrier ?FS. g(f x) = x" and fg: "\y \ carrier ?FT. f(g y) = y" by (auto simp: group_isomorphisms_def) let ?f = "\h. restrict (h \ g) (carrier ?FT)" let ?g = "\h. restrict (h \ f) (carrier ?FS)" show ?thesis proof (rule lepoll_antisym) show "?HS \ ?HT" unfolding lepoll_def proof (intro exI conjI) show "inj_on ?f ?HS" apply (rule inj_on_inverseI [where g = ?g]) using hom_in_carrier [OF f] by (auto simp: gf fun_eq_iff carrier_integer_mod_group Ball_def Pi_def extensional_def) show "?f ` ?HS \ ?HT" proof clarsimp fix h assume h: "h \ hom ?FS (integer_mod_group 2)" have "h \ g \ hom ?FT (integer_mod_group 2)" by (rule hom_compose [OF g h]) moreover have "restrict (h \ g) (carrier ?FT) x = (h \ g) x" if "x \ carrier ?FT" for x using g that by (simp add: hom_def) ultimately show "restrict (h \ g) (carrier ?FT) \ hom ?FT (integer_mod_group 2)" using T.hom_restrict by fastforce qed qed next show "?HT \ ?HS" unfolding lepoll_def proof (intro exI conjI) show "inj_on ?g ?HT" apply (rule inj_on_inverseI [where g = ?f]) using hom_in_carrier [OF g] by (auto simp: fg fun_eq_iff carrier_integer_mod_group Ball_def Pi_def extensional_def) show "?g ` ?HT \ ?HS" proof clarsimp fix k assume k: "k \ hom ?FT (integer_mod_group 2)" have "k \ f \ hom ?FS (integer_mod_group 2)" by (rule hom_compose [OF f k]) moreover have "restrict (k \ f) (carrier ?FS) x = (k \ f) x" if "x \ carrier ?FS" for x using f that by (simp add: hom_def) ultimately show "restrict (k \ f) (carrier ?FS) \ hom ?FS (integer_mod_group 2)" using S.hom_restrict by fastforce qed qed qed qed also have "\ \ T \\<^sub>E ?Two" using G2.eqpoll_homomorphisms_from_free_Abelian_group by (simp add: carrier_integer_mod_group) finally have *: "S \\<^sub>E ?Two \ T \\<^sub>E ?Two" . then have "finite (S \\<^sub>E ?Two) \ finite (T \\<^sub>E ?Two)" by (rule eqpoll_finite_iff) then have "finite S \ finite T" by (auto simp: finite_funcset_iff) then consider "finite S" "finite T" | "~ finite S" "~ finite T" by blast then show ?rhs proof cases case 1 with * have "2 ^ card S = (2::nat) ^ card T" by (simp add: card_PiE finite_PiE eqpoll_iff_card) then have "card S = card T" by auto then show ?thesis using eqpoll_iff_card 1 by blast next case 2 have "carrier (free_Abelian_group S) \ carrier (free_Abelian_group T)" using L by (simp add: iso_imp_eqpoll_carrier) then show ?thesis using 2 eqpoll_free_Abelian_group_infinite eqpoll_sym eqpoll_trans by metis qed next assume ?rhs then obtain f g where f: "\x. x \ S \ f x \ T \ g(f x) = x" and g: "\y. y \ T \ g y \ S \ f(g y) = y" using eqpoll_iff_bijections by metis interpret S: comm_group "?FS" by (simp add: abelian_free_Abelian_group) interpret T: comm_group "?FT" by (simp add: abelian_free_Abelian_group) have "(frag_of \ f) ` S \ carrier (free_Abelian_group T)" using f by auto then obtain h where h: "h \ hom (free_Abelian_group S) (free_Abelian_group T)" and h_frag: "\x. x \ S \ h (frag_of x) = (frag_of \ f) x" using T.free_Abelian_group_universal [of "frag_of \ f" S] by blast interpret hhom: group_hom "free_Abelian_group S" "free_Abelian_group T" h by (simp add: h group_hom_axioms_def group_hom_def) have "(frag_of \ g) ` T \ carrier (free_Abelian_group S)" using g by auto then obtain k where k: "k \ hom (free_Abelian_group T) (free_Abelian_group S)" and k_frag: "\x. x \ T \ k (frag_of x) = (frag_of \ g) x" using S.free_Abelian_group_universal [of "frag_of \ g" T] by blast interpret khom: group_hom "free_Abelian_group T" "free_Abelian_group S" k by (simp add: k group_hom_axioms_def group_hom_def) have kh: "Poly_Mapping.keys x \ S \ Poly_Mapping.keys x \ S \ k (h x) = x" for x proof (induction rule: frag_induction) case zero then show ?case apply auto by (metis group_free_Abelian_group h hom_one k one_free_Abelian_group) next case (one x) then show ?case by (auto simp: h_frag k_frag f) next case (diff a b) with keys_diff have "Poly_Mapping.keys (a - b) \ S" by (metis Un_least order_trans) with diff hhom.hom_closed show ?case by (simp add: hom_frag_diff [OF h] hom_frag_diff [OF k]) qed have hk: "Poly_Mapping.keys y \ T \ Poly_Mapping.keys y \ T \ h (k y) = y" for y proof (induction rule: frag_induction) case zero then show ?case apply auto by (metis group_free_Abelian_group h hom_one k one_free_Abelian_group) next case (one y) then show ?case by (auto simp: h_frag k_frag g) next case (diff a b) with keys_diff have "Poly_Mapping.keys (a - b) \ T" by (metis Un_least order_trans) with diff khom.hom_closed show ?case by (simp add: hom_frag_diff [OF h] hom_frag_diff [OF k]) qed have "h \ iso ?FS ?FT" unfolding iso_def bij_betw_iff_bijections mem_Collect_eq proof (intro conjI exI ballI h) fix x assume x: "x \ carrier (free_Abelian_group S)" show "h x \ carrier (free_Abelian_group T)" by (meson x h hom_in_carrier) show "k (h x) = x" using x by (simp add: kh) next fix y assume y: "y \ carrier (free_Abelian_group T)" show "k y \ carrier (free_Abelian_group S)" by (meson y k hom_in_carrier) show "h (k y) = y" using y by (simp add: hk) qed then show "?FS \ ?FT" by (auto simp: is_iso_def) qed lemma isomorphic_group_integer_free_Abelian_group_singleton: "integer_group \ free_Abelian_group {x}" proof - have "(\n. frag_cmul n (frag_of x)) \ iso integer_group (free_Abelian_group {x})" proof (rule isoI [OF homI]) show "bij_betw (\n. frag_cmul n (frag_of x)) (carrier integer_group) (carrier (free_Abelian_group {x}))" apply (rule bij_betwI [where g = "\y. Poly_Mapping.lookup y x"]) by (auto simp: integer_group_def in_keys_iff intro!: poly_mapping_eqI) qed (auto simp: frag_cmul_distrib) then show ?thesis unfolding is_iso_def by blast qed lemma group_hom_free_Abelian_groups_id: "id \ hom (free_Abelian_group S) (free_Abelian_group T) \ S \ T" proof - have "x \ T" if ST: "\c:: 'a \\<^sub>0 int. Poly_Mapping.keys c \ S \ Poly_Mapping.keys c \ T" and "x \ S" for x using ST [of "frag_of x"] \x \ S\ by simp then show ?thesis by (auto simp: hom_def free_Abelian_group_def Pi_def) qed proposition iso_free_Abelian_group_sum: assumes "pairwise (\i j. disjnt (S i) (S j)) I" shows "(\f. sum' f I) \ iso (sum_group I (\i. free_Abelian_group(S i))) (free_Abelian_group (\(S ` I)))" (is "?h \ iso ?G ?H") proof (rule isoI) show hom: "?h \ hom ?G ?H" proof (rule homI) show "?h c \ carrier ?H" if "c \ carrier ?G" for c using that apply (simp add: sum.G_def carrier_sum_group) apply (rule order_trans [OF keys_sum]) apply (auto simp: free_Abelian_group_def) done show "?h (x \\<^bsub>?G\<^esub> y) = ?h x \\<^bsub>?H\<^esub> ?h y" if "x \ carrier ?G" "y \ carrier ?G" for x y using that by (simp add: sum.finite_Collect_op carrier_sum_group sum.distrib') qed interpret GH: group_hom "?G" "?H" "?h" using hom by (simp add: group_hom_def group_hom_axioms_def) show "bij_betw ?h (carrier ?G) (carrier ?H)" unfolding bij_betw_def proof (intro conjI subset_antisym) show "?h ` carrier ?G \ carrier ?H" apply (clarsimp simp: sum.G_def carrier_sum_group simp del: carrier_free_Abelian_group_iff) by (force simp: PiE_def Pi_iff intro!: sum_closed_free_Abelian_group) have *: "poly_mapping.lookup (Abs_poly_mapping (\j. if j \ S i then poly_mapping.lookup x j else 0)) k = (if k \ S i then poly_mapping.lookup x k else 0)" if "i \ I" for i k and x :: "'b \\<^sub>0 int" using that by (auto simp: conj_commute cong: conj_cong) have eq: "Abs_poly_mapping (\j. if j \ S i then poly_mapping.lookup x j else 0) = 0 \ (\c \ S i. poly_mapping.lookup x c = 0)" if "i \ I" for i and x :: "'b \\<^sub>0 int" apply (auto simp: poly_mapping_eq_iff fun_eq_iff) apply (simp add: * Abs_poly_mapping_inverse conj_commute cong: conj_cong) apply (force dest!: spec split: if_split_asm) done have "x \ ?h ` {x \ \\<^sub>E i\I. {c. Poly_Mapping.keys c \ S i}. finite {i \ I. x i \ 0}}" if x: "Poly_Mapping.keys x \ \ (S ` I)" for x :: "'b \\<^sub>0 int" proof - let ?f = "(\i c. if c \ S i then Poly_Mapping.lookup x c else 0)" define J where "J \ {i \ I. \c\S i. c \ Poly_Mapping.keys x}" have "J \ (\c. THE i. i \ I \ c \ S i) ` Poly_Mapping.keys x" proof (clarsimp simp: J_def) show "i \ (\c. THE i. i \ I \ c \ S i) ` Poly_Mapping.keys x" if "i \ I" "c \ S i" "c \ Poly_Mapping.keys x" for i c proof show "i = (THE i. i \ I \ c \ S i)" using assms that by (auto simp: pairwise_def disjnt_def intro: the_equality [symmetric]) qed (simp add: that) qed then have fin: "finite J" using finite_subset finite_keys by blast have [simp]: "Poly_Mapping.keys (Abs_poly_mapping (?f i)) = {k. ?f i k \ 0}" if "i \ I" for i by (simp add: eq_onp_def keys.abs_eq conj_commute cong: conj_cong) have [simp]: "Poly_Mapping.lookup (Abs_poly_mapping (?f i)) c = ?f i c" if "i \ I" for i c by (auto simp: Abs_poly_mapping_inverse conj_commute cong: conj_cong) show ?thesis proof have "poly_mapping.lookup x c = poly_mapping.lookup (?h (\i\I. Abs_poly_mapping (?f i))) c" for c proof (cases "c \ Poly_Mapping.keys x") case True then obtain i where "i \ I" "c \ S i" "?f i c \ 0" using x by (auto simp: in_keys_iff) then have 1: "poly_mapping.lookup (sum' (\j. Abs_poly_mapping (?f j)) (I - {i})) c = 0" using assms apply (simp add: sum.G_def Poly_Mapping.lookup_sum pairwise_def disjnt_def) apply (force simp: eq split: if_split_asm intro!: comm_monoid_add_class.sum.neutral) done have 2: "poly_mapping.lookup x c = poly_mapping.lookup (Abs_poly_mapping (?f i)) c" by (auto simp: \c \ S i\ Abs_poly_mapping_inverse conj_commute cong: conj_cong) have "finite {i \ I. Abs_poly_mapping (?f i) \ 0}" by (rule finite_subset [OF _ fin]) (use \i \ I\ J_def eq in \auto simp: in_keys_iff\) with \i \ I\ have "?h (\j\I. Abs_poly_mapping (?f j)) = Abs_poly_mapping (?f i) + sum' (\j. Abs_poly_mapping (?f j)) (I - {i})" by (simp add: sum_diff1') then show ?thesis by (simp add: 1 2 Poly_Mapping.lookup_add) next case False then have "poly_mapping.lookup x c = 0" using keys.rep_eq by force then show ?thesis unfolding sum.G_def by (simp add: lookup_sum * comm_monoid_add_class.sum.neutral) qed then show "x = ?h (\i\I. Abs_poly_mapping (?f i))" by (rule poly_mapping_eqI) have "(\i. Abs_poly_mapping (?f i)) \ (\ i\I. {c. Poly_Mapping.keys c \ S i})" by (auto simp: PiE_def Pi_def in_keys_iff) then show "(\i\I. Abs_poly_mapping (?f i)) \ {x \ \\<^sub>E i\I. {c. Poly_Mapping.keys c \ S i}. finite {i \ I. x i \ 0}}" - using fin unfolding J_def by (simp add: eq in_keys_iff cong: conj_cong) + using fin unfolding J_def by (force simp add: eq in_keys_iff cong: conj_cong) qed qed then show "carrier ?H \ ?h ` carrier ?G" by (simp add: carrier_sum_group) (auto simp: free_Abelian_group_def) show "inj_on ?h (carrier (sum_group I (\i. free_Abelian_group (S i))))" unfolding GH.inj_on_one_iff proof clarify fix x assume "x \ carrier ?G" "?h x = \\<^bsub>?H\<^esub>" then have eq0: "sum' x I = 0" and xs: "\i. i \ I \ Poly_Mapping.keys (x i) \ S i" and xext: "x \ extensional I" and fin: "finite {i \ I. x i \ 0}" by (simp_all add: carrier_sum_group PiE_def Pi_def) have "x i = 0" if "i \ I" for i proof - have "sum' x (insert i (I - {i})) = 0" using eq0 that by (simp add: insert_absorb) moreover have "Poly_Mapping.keys (sum' x (I - {i})) = {}" proof - have "x i = - sum' x (I - {i})" by (metis (mono_tags, lifting) diff_zero eq0 fin sum_diff1' minus_diff_eq that) then have "Poly_Mapping.keys (x i) = Poly_Mapping.keys (sum' x (I - {i}))" by simp then have "Poly_Mapping.keys (sum' x (I - {i})) \ S i" using that xs by metis moreover have "Poly_Mapping.keys (sum' x (I - {i})) \ (\j \ I - {i}. S j)" proof - have "Poly_Mapping.keys (sum' x (I - {i})) \ (\i\{j \ I. j \ i \ x j \ 0}. Poly_Mapping.keys (x i))" using keys_sum [of x "{j \ I. j \ i \ x j \ 0}"] by (simp add: sum.G_def) also have "\ \ \ (S ` (I - {i}))" using xs by force finally show ?thesis . qed moreover have "A = {}" if "A \ S i" "A \ \ (S ` (I - {i}))" for A using assms that \i \ I\ by (force simp: pairwise_def disjnt_def image_def subset_iff) ultimately show ?thesis by metis qed then have [simp]: "sum' x (I - {i}) = 0" by (auto simp: sum.G_def) have "sum' x (insert i (I - {i})) = x i" by (subst sum.insert' [OF finite_subset [OF _ fin]]) auto ultimately show ?thesis by metis qed with xext [unfolded extensional_def] show "x = \\<^bsub>sum_group I (\i. free_Abelian_group (S i))\<^esub>" by (force simp: free_Abelian_group_def) qed qed qed lemma isomorphic_free_Abelian_group_Union: "pairwise disjnt I \ free_Abelian_group(\ I) \ sum_group I free_Abelian_group" using iso_free_Abelian_group_sum [of "\X. X" I] by (metis SUP_identity_eq empty_iff group.iso_sym group_free_Abelian_group is_iso_def sum_group) lemma isomorphic_sum_integer_group: "sum_group I (\i. integer_group) \ free_Abelian_group I" proof - have "sum_group I (\i. integer_group) \ sum_group I (\i. free_Abelian_group {i})" by (rule iso_sum_groupI) (auto simp: isomorphic_group_integer_free_Abelian_group_singleton) also have "\ \ free_Abelian_group I" using iso_free_Abelian_group_sum [of "\x. {x}" I] by (auto simp: is_iso_def) finally show ?thesis . qed end diff --git a/src/HOL/Library/FuncSet.thy b/src/HOL/Library/FuncSet.thy --- a/src/HOL/Library/FuncSet.thy +++ b/src/HOL/Library/FuncSet.thy @@ -1,717 +1,717 @@ (* Title: HOL/Library/FuncSet.thy Author: Florian Kammueller and Lawrence C Paulson, Lukas Bulwahn *) section \Pi and Function Sets\ theory FuncSet imports Main abbrevs PiE = "Pi\<^sub>E" and PIE = "\\<^sub>E" begin definition Pi :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" where "Pi A B = {f. \x. x \ A \ f x \ B x}" definition extensional :: "'a set \ ('a \ 'b) set" where "extensional A = {f. \x. x \ A \ f x = undefined}" definition "restrict" :: "('a \ 'b) \ 'a set \ 'a \ 'b" where "restrict f A = (\x. if x \ A then f x else undefined)" abbreviation funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 60) where "A \ B \ Pi A (\_. B)" syntax "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\ _\_./ _)" 10) "_lam" :: "pttrn \ 'a set \ ('a \ 'b) \ ('a \ 'b)" ("(3\_\_./ _)" [0,0,3] 3) translations "\ x\A. B" \ "CONST Pi A (\x. B)" "\x\A. f" \ "CONST restrict (\x. f) A" definition "compose" :: "'a set \ ('b \ 'c) \ ('a \ 'b) \ ('a \ 'c)" where "compose A g f = (\x\A. g (f x))" subsection \Basic Properties of \<^term>\Pi\\ lemma Pi_I[intro!]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" by (simp add: Pi_def) lemma Pi_I'[simp]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" by (simp add:Pi_def) lemma funcsetI: "(\x. x \ A \ f x \ B) \ f \ A \ B" by (simp add: Pi_def) lemma Pi_mem: "f \ Pi A B \ x \ A \ f x \ B x" by (simp add: Pi_def) lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" unfolding Pi_def by auto lemma PiE [elim]: "f \ Pi A B \ (f x \ B x \ Q) \ (x \ A \ Q) \ Q" by (auto simp: Pi_def) lemma Pi_cong: "(\w. w \ A \ f w = g w) \ f \ Pi A B \ g \ Pi A B" by (auto simp: Pi_def) lemma funcset_id [simp]: "(\x. x) \ A \ A" by auto lemma funcset_mem: "f \ A \ B \ x \ A \ f x \ B" by (simp add: Pi_def) lemma funcset_image: "f \ A \ B \ f ` A \ B" by auto lemma image_subset_iff_funcset: "F ` A \ B \ F \ A \ B" by auto lemma funcset_to_empty_iff: "A \ {} = (if A={} then UNIV else {})" by auto lemma Pi_eq_empty[simp]: "(\ x \ A. B x) = {} \ (\x\A. B x = {})" proof - - have "\x\A. B x = {}" if "\f. \y. y \ A \ f y \ B y" + have "\x\A. B x = {}" if "\f. \y. y \ A \ f y \ B y" using that [of "\u. SOME y. y \ B u"] some_in_eq by blast then show ?thesis by force qed lemma Pi_empty [simp]: "Pi {} B = UNIV" by (simp add: Pi_def) lemma Pi_Int: "Pi I E \ Pi I F = (\ i\I. E i \ F i)" by auto lemma Pi_UN: fixes A :: "nat \ 'i \ 'a set" assumes "finite I" and mono: "\i n m. i \ I \ n \ m \ A n i \ A m i" shows "(\n. Pi I (A n)) = (\ i\I. \n. A n i)" proof (intro set_eqI iffI) fix f assume "f \ (\ i\I. \n. A n i)" then have "\i\I. \n. f i \ A n i" by auto from bchoice[OF this] obtain n where n: "f i \ A (n i) i" if "i \ I" for i by auto obtain k where k: "n i \ k" if "i \ I" for i using \finite I\ finite_nat_set_iff_bounded_le[of "n`I"] by auto have "f \ Pi I (A k)" proof (intro Pi_I) fix i assume "i \ I" from mono[OF this, of "n i" k] k[OF this] n[OF this] show "f i \ A k i" by auto qed then show "f \ (\n. Pi I (A n))" by auto qed auto lemma Pi_UNIV [simp]: "A \ UNIV = UNIV" by (simp add: Pi_def) text \Covariance of Pi-sets in their second argument\ lemma Pi_mono: "(\x. x \ A \ B x \ C x) \ Pi A B \ Pi A C" by auto text \Contravariance of Pi-sets in their first argument\ lemma Pi_anti_mono: "A' \ A \ Pi A B \ Pi A' B" by auto lemma prod_final: assumes 1: "fst \ f \ Pi A B" and 2: "snd \ f \ Pi A C" shows "f \ (\ z \ A. B z \ C z)" proof (rule Pi_I) fix z assume z: "z \ A" have "f z = (fst (f z), snd (f z))" by simp also have "\ \ B z \ C z" by (metis SigmaI PiE o_apply 1 2 z) finally show "f z \ B z \ C z" . qed lemma Pi_split_domain[simp]: "x \ Pi (I \ J) X \ x \ Pi I X \ x \ Pi J X" by (auto simp: Pi_def) lemma Pi_split_insert_domain[simp]: "x \ Pi (insert i I) X \ x \ Pi I X \ x i \ X i" by (auto simp: Pi_def) lemma Pi_cancel_fupd_range[simp]: "i \ I \ x \ Pi I (B(i := b)) \ x \ Pi I B" by (auto simp: Pi_def) lemma Pi_cancel_fupd[simp]: "i \ I \ x(i := a) \ Pi I B \ x \ Pi I B" by (auto simp: Pi_def) lemma Pi_fupd_iff: "i \ I \ f \ Pi I (B(i := A)) \ f \ Pi (I - {i}) B \ f i \ A" apply auto apply (metis PiE fun_upd_apply) by force subsection \Composition With a Restricted Domain: \<^term>\compose\\ lemma funcset_compose: "f \ A \ B \ g \ B \ C \ compose A g f \ A \ C" by (simp add: Pi_def compose_def restrict_def) lemma compose_assoc: assumes "f \ A \ B" shows "compose A h (compose A g f) = compose A (compose B h g) f" using assms by (simp add: fun_eq_iff Pi_def compose_def restrict_def) lemma compose_eq: "x \ A \ compose A g f x = g (f x)" by (simp add: compose_def restrict_def) lemma surj_compose: "f ` A = B \ g ` B = C \ compose A g f ` A = C" by (auto simp add: image_def compose_eq) subsection \Bounded Abstraction: \<^term>\restrict\\ lemma restrict_cong: "I = J \ (\i. i \ J =simp=> f i = g i) \ restrict f I = restrict g J" by (auto simp: restrict_def fun_eq_iff simp_implies_def) lemma restrictI[intro!]: "(\x. x \ A \ f x \ B x) \ (\x\A. f x) \ Pi A B" by (simp add: Pi_def restrict_def) lemma restrict_apply[simp]: "(\y\A. f y) x = (if x \ A then f x else undefined)" by (simp add: restrict_def) lemma restrict_apply': "x \ A \ (\y\A. f y) x = f x" by simp lemma restrict_ext: "(\x. x \ A \ f x = g x) \ (\x\A. f x) = (\x\A. g x)" by (simp add: fun_eq_iff Pi_def restrict_def) lemma restrict_UNIV: "restrict f UNIV = f" by (simp add: restrict_def) lemma inj_on_restrict_eq [simp]: "inj_on (restrict f A) A = inj_on f A" by (simp add: inj_on_def restrict_def) lemma Id_compose: "f \ A \ B \ f \ extensional A \ compose A (\y\B. y) f = f" by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) lemma compose_Id: "g \ A \ B \ g \ extensional A \ compose A g (\x\A. x) = g" by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) lemma image_restrict_eq [simp]: "(restrict f A) ` A = f ` A" by (auto simp add: restrict_def) lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \ B)" unfolding restrict_def by (simp add: fun_eq_iff) lemma restrict_fupd[simp]: "i \ I \ restrict (f (i := x)) I = restrict f I" by (auto simp: restrict_def) lemma restrict_upd[simp]: "i \ I \ (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" by (auto simp: fun_eq_iff) lemma restrict_Pi_cancel: "restrict x I \ Pi I A \ x \ Pi I A" by (auto simp: restrict_def Pi_def) lemma sum_restrict' [simp]: "sum' (\i\I. g i) I = sum' (\i. g i) I" by (simp add: sum.G_def conj_commute cong: conj_cong) lemma prod_restrict' [simp]: "prod' (\i\I. g i) I = prod' (\i. g i) I" by (simp add: prod.G_def conj_commute cong: conj_cong) subsection \Bijections Between Sets\ text \The definition of \<^const>\bij_betw\ is in \Fun.thy\, but most of the theorems belong here, or need at least \<^term>\Hilbert_Choice\.\ lemma bij_betwI: assumes "f \ A \ B" and "g \ B \ A" and g_f: "\x. x\A \ g (f x) = x" and f_g: "\y. y\B \ f (g y) = y" shows "bij_betw f A B" unfolding bij_betw_def proof show "inj_on f A" by (metis g_f inj_on_def) have "f ` A \ B" using \f \ A \ B\ by auto moreover have "B \ f ` A" by auto (metis Pi_mem \g \ B \ A\ f_g image_iff) ultimately show "f ` A = B" by blast qed lemma bij_betw_imp_funcset: "bij_betw f A B \ f \ A \ B" by (auto simp add: bij_betw_def) lemma inj_on_compose: "bij_betw f A B \ inj_on g B \ inj_on (compose A g f) A" by (auto simp add: bij_betw_def inj_on_def compose_eq) lemma bij_betw_compose: "bij_betw f A B \ bij_betw g B C \ bij_betw (compose A g f) A C" apply (simp add: bij_betw_def compose_eq inj_on_compose) apply (auto simp add: compose_def image_def) done lemma bij_betw_restrict_eq [simp]: "bij_betw (restrict f A) A B = bij_betw f A B" by (simp add: bij_betw_def) subsection \Extensionality\ lemma extensional_empty[simp]: "extensional {} = {\x. undefined}" unfolding extensional_def by auto lemma extensional_arb: "f \ extensional A \ x \ A \ f x = undefined" by (simp add: extensional_def) lemma restrict_extensional [simp]: "restrict f A \ extensional A" by (simp add: restrict_def extensional_def) lemma compose_extensional [simp]: "compose A f g \ extensional A" by (simp add: compose_def) lemma extensionalityI: assumes "f \ extensional A" and "g \ extensional A" and "\x. x \ A \ f x = g x" shows "f = g" using assms by (force simp add: fun_eq_iff extensional_def) lemma extensional_restrict: "f \ extensional A \ restrict f A = f" by (rule extensionalityI[OF restrict_extensional]) auto lemma extensional_subset: "f \ extensional A \ A \ B \ f \ extensional B" unfolding extensional_def by auto lemma inv_into_funcset: "f ` A = B \ (\x\B. inv_into A f x) \ B \ A" by (unfold inv_into_def) (fast intro: someI2) lemma compose_inv_into_id: "bij_betw f A B \ compose A (\y\B. inv_into A f y) f = (\x\A. x)" apply (simp add: bij_betw_def compose_def) apply (rule restrict_ext, auto) done lemma compose_id_inv_into: "f ` A = B \ compose B f (\y\B. inv_into A f y) = (\x\B. x)" apply (simp add: compose_def) apply (rule restrict_ext) apply (simp add: f_inv_into_f) done lemma extensional_insert[intro, simp]: assumes "a \ extensional (insert i I)" shows "a(i := b) \ extensional (insert i I)" using assms unfolding extensional_def by auto lemma extensional_Int[simp]: "extensional I \ extensional I' = extensional (I \ I')" unfolding extensional_def by auto lemma extensional_UNIV[simp]: "extensional UNIV = UNIV" by (auto simp: extensional_def) lemma restrict_extensional_sub[intro]: "A \ B \ restrict f A \ extensional B" unfolding restrict_def extensional_def by auto lemma extensional_insert_undefined[intro, simp]: "a \ extensional (insert i I) \ a(i := undefined) \ extensional I" unfolding extensional_def by auto lemma extensional_insert_cancel[intro, simp]: "a \ extensional I \ a \ extensional (insert i I)" unfolding extensional_def by auto subsection \Cardinality\ lemma card_inj: "f \ A \ B \ inj_on f A \ finite B \ card A \ card B" by (rule card_inj_on_le) auto lemma card_bij: assumes "f \ A \ B" "inj_on f A" and "g \ B \ A" "inj_on g B" and "finite A" "finite B" shows "card A = card B" using assms by (blast intro: card_inj order_antisym) subsection \Extensional Function Spaces\ definition PiE :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" where "PiE S T = Pi S T \ extensional S" abbreviation "Pi\<^sub>E A B \ PiE A B" syntax "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\\<^sub>E _\_./ _)" 10) translations "\\<^sub>E x\A. B" \ "CONST Pi\<^sub>E A (\x. B)" abbreviation extensional_funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\\<^sub>E" 60) where "A \\<^sub>E B \ (\\<^sub>E i\A. B)" lemma extensional_funcset_def: "extensional_funcset S T = (S \ T) \ extensional S" by (simp add: PiE_def) lemma PiE_empty_domain[simp]: "Pi\<^sub>E {} T = {\x. undefined}" unfolding PiE_def by simp lemma PiE_UNIV_domain: "Pi\<^sub>E UNIV T = Pi UNIV T" unfolding PiE_def by simp lemma PiE_empty_range[simp]: "i \ I \ F i = {} \ (\\<^sub>E i\I. F i) = {}" unfolding PiE_def by auto lemma PiE_eq_empty_iff: "Pi\<^sub>E I F = {} \ (\i\I. F i = {})" proof assume "Pi\<^sub>E I F = {}" show "\i\I. F i = {}" proof (rule ccontr) assume "\ ?thesis" then have "\i. \y. (i \ I \ y \ F i) \ (i \ I \ y = undefined)" by auto from choice[OF this] obtain f where " \x. (x \ I \ f x \ F x) \ (x \ I \ f x = undefined)" .. then have "f \ Pi\<^sub>E I F" by (auto simp: extensional_def PiE_def) with \Pi\<^sub>E I F = {}\ show False by auto qed qed (auto simp: PiE_def) lemma PiE_arb: "f \ Pi\<^sub>E S T \ x \ S \ f x = undefined" unfolding PiE_def by auto (auto dest!: extensional_arb) lemma PiE_mem: "f \ Pi\<^sub>E S T \ x \ S \ f x \ T x" unfolding PiE_def by auto lemma PiE_fun_upd: "y \ T x \ f \ Pi\<^sub>E S T \ f(x := y) \ Pi\<^sub>E (insert x S) T" unfolding PiE_def extensional_def by auto lemma fun_upd_in_PiE: "x \ S \ f \ Pi\<^sub>E (insert x S) T \ f(x := undefined) \ Pi\<^sub>E S T" unfolding PiE_def extensional_def by auto lemma PiE_insert_eq: "Pi\<^sub>E (insert x S) T = (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" proof - { fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" by (auto intro!: image_eqI[where x="(f x, f(x := undefined))"] intro: fun_upd_in_PiE PiE_mem) } moreover { fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" by (auto intro!: image_eqI[where x="(f x, f)"] intro: fun_upd_in_PiE PiE_mem simp: insert_absorb) } ultimately show ?thesis by (auto intro: PiE_fun_upd) qed lemma PiE_Int: "Pi\<^sub>E I A \ Pi\<^sub>E I B = Pi\<^sub>E I (\x. A x \ B x)" by (auto simp: PiE_def) lemma PiE_cong: "(\i. i\I \ A i = B i) \ Pi\<^sub>E I A = Pi\<^sub>E I B" unfolding PiE_def by (auto simp: Pi_cong) lemma PiE_E [elim]: assumes "f \ Pi\<^sub>E A B" obtains "x \ A" and "f x \ B x" | "x \ A" and "f x = undefined" using assms by (auto simp: Pi_def PiE_def extensional_def) lemma PiE_I[intro!]: "(\x. x \ A \ f x \ B x) \ (\x. x \ A \ f x = undefined) \ f \ Pi\<^sub>E A B" by (simp add: PiE_def extensional_def) lemma PiE_mono: "(\x. x \ A \ B x \ C x) \ Pi\<^sub>E A B \ Pi\<^sub>E A C" by auto lemma PiE_iff: "f \ Pi\<^sub>E I X \ (\i\I. f i \ X i) \ f \ extensional I" by (simp add: PiE_def Pi_iff) -lemma restrict_PiE_iff [simp]: "restrict f I \ Pi\<^sub>E I X \ (\i \ I. f i \ X i)" +lemma restrict_PiE_iff: "restrict f I \ Pi\<^sub>E I X \ (\i \ I. f i \ X i)" by (simp add: PiE_iff) lemma ext_funcset_to_sing_iff [simp]: "A \\<^sub>E {a} = {\x\A. a}" by (auto simp: PiE_def Pi_iff extensionalityI) lemma PiE_restrict[simp]: "f \ Pi\<^sub>E A B \ restrict f A = f" by (simp add: extensional_restrict PiE_def) lemma restrict_PiE[simp]: "restrict f I \ Pi\<^sub>E I S \ f \ Pi I S" by (auto simp: PiE_iff) lemma PiE_eq_subset: assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" and eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" and "i \ I" shows "F i \ F' i" proof fix x assume "x \ F i" with ne have "\j. \y. (j \ I \ y \ F j \ (i = j \ x = y)) \ (j \ I \ y = undefined)" by auto from choice[OF this] obtain f where f: " \j. (j \ I \ f j \ F j \ (i = j \ x = f j)) \ (j \ I \ f j = undefined)" .. then have "f \ Pi\<^sub>E I F" by (auto simp: extensional_def PiE_def) then have "f \ Pi\<^sub>E I F'" using assms by simp then show "x \ F' i" using f \i \ I\ by (auto simp: PiE_def) qed lemma PiE_eq_iff_not_empty: assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" shows "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i)" proof (intro iffI ballI) fix i assume eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" assume i: "i \ I" show "F i = F' i" using PiE_eq_subset[of I F F', OF ne eq i] using PiE_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] by auto qed (auto simp: PiE_def) lemma PiE_eq_iff: "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i) \ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" proof (intro iffI disjCI) assume eq[simp]: "Pi\<^sub>E I F = Pi\<^sub>E I F'" assume "\ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" then have "(\i\I. F i \ {}) \ (\i\I. F' i \ {})" using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by auto with PiE_eq_iff_not_empty[of I F F'] show "\i\I. F i = F' i" by auto next assume "(\i\I. F i = F' i) \ (\i\I. F i = {}) \ (\i\I. F' i = {})" then show "Pi\<^sub>E I F = Pi\<^sub>E I F'" using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by (auto simp: PiE_def) qed lemma extensional_funcset_fun_upd_restricts_rangeI: "\y \ S. f x \ f y \ f \ (insert x S) \\<^sub>E T \ f(x := undefined) \ S \\<^sub>E (T - {f x})" unfolding extensional_funcset_def extensional_def by (auto split: if_split_asm) lemma extensional_funcset_fun_upd_extends_rangeI: assumes "a \ T" "f \ S \\<^sub>E (T - {a})" shows "f(x := a) \ insert x S \\<^sub>E T" using assms unfolding extensional_funcset_def extensional_def by auto lemma subset_PiE: "PiE I S \ PiE I T \ PiE I S = {} \ (\i \ I. S i \ T i)" (is "?lhs \ _ \ ?rhs") proof (cases "PiE I S = {}") case False moreover have "?lhs = ?rhs" proof assume L: ?lhs have "\i. i\I \ S i \ {}" using False PiE_eq_empty_iff by blast with L show ?rhs by (simp add: PiE_Int PiE_eq_iff inf.absorb_iff2) qed auto ultimately show ?thesis by simp qed simp lemma PiE_eq: "PiE I S = PiE I T \ PiE I S = {} \ PiE I T = {} \ (\i \ I. S i = T i)" by (auto simp: PiE_eq_iff PiE_eq_empty_iff) lemma PiE_UNIV [simp]: "PiE UNIV (\i. UNIV) = UNIV" by blast lemma image_projection_PiE: "(\f. f i) ` (PiE I S) = (if PiE I S = {} then {} else if i \ I then S i else {undefined})" proof - have "(\f. f i) ` Pi\<^sub>E I S = S i" if "i \ I" "f \ PiE I S" for f using that apply auto by (rule_tac x="(\k. if k=i then x else f k)" in image_eqI) auto moreover have "(\f. f i) ` Pi\<^sub>E I S = {undefined}" if "f \ PiE I S" "i \ I" for f using that by (blast intro: PiE_arb [OF that, symmetric]) ultimately show ?thesis by auto qed -lemma PiE_singleton: +lemma PiE_singleton: assumes "f \ extensional A" shows "PiE A (\x. {f x}) = {f}" proof - { fix g assume "g \ PiE A (\x. {f x})" hence "g x = f x" for x using assms by (cases "x \ A") (auto simp: extensional_def) hence "g = f" by (simp add: fun_eq_iff) } thus ?thesis using assms by (auto simp: extensional_def) qed lemma PiE_eq_singleton: "(\\<^sub>E i\I. S i) = {\i\I. f i} \ (\i\I. S i = {f i})" by (metis (mono_tags, lifting) PiE_eq PiE_singleton insert_not_empty restrict_apply' restrict_extensional) lemma PiE_over_singleton_iff: "(\\<^sub>E x\{a}. B x) = (\b \ B a. {\x \ {a}. b})" apply (auto simp: PiE_iff split: if_split_asm) apply (metis (no_types, lifting) extensionalityI restrict_apply' restrict_extensional singletonD) done lemma all_PiE_elements: "(\z \ PiE I S. \i \ I. P i (z i)) \ PiE I S = {} \ (\i \ I. \x \ S i. P i x)" (is "?lhs = ?rhs") proof (cases "PiE I S = {}") case False then obtain f where f: "\i. i \ I \ f i \ S i" by fastforce show ?thesis proof assume L: ?lhs have "P i x" if "i \ I" "x \ S i" for i x proof - have "(\j \ I. if j=i then x else f j) \ PiE I S" by (simp add: f that(2)) then have "P i ((\j \ I. if j=i then x else f j) i)" using L that(1) by blast with that show ?thesis by simp qed then show ?rhs by (simp add: False) qed fastforce qed simp lemma PiE_ext: "\x \ PiE k s; y \ PiE k s; \i. i \ k \ x i = y i\ \ x = y" by (metis ext PiE_E) subsubsection \Injective Extensional Function Spaces\ lemma extensional_funcset_fun_upd_inj_onI: assumes "f \ S \\<^sub>E (T - {a})" and "inj_on f S" shows "inj_on (f(x := a)) S" using assms unfolding extensional_funcset_def by (auto intro!: inj_on_fun_updI) lemma extensional_funcset_extend_domain_inj_on_eq: assumes "x \ S" shows "{f. f \ (insert x S) \\<^sub>E T \ inj_on f (insert x S)} = (\(y, g). g(x:=y)) ` {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" using assms apply (auto del: PiE_I PiE_E) apply (auto intro: extensional_funcset_fun_upd_inj_onI extensional_funcset_fun_upd_extends_rangeI del: PiE_I PiE_E) apply (auto simp add: image_iff inj_on_def) apply (rule_tac x="xa x" in exI) apply (auto intro: PiE_mem del: PiE_I PiE_E) apply (rule_tac x="xa(x := undefined)" in exI) apply (auto intro!: extensional_funcset_fun_upd_restricts_rangeI) apply (auto dest!: PiE_mem split: if_split_asm) done lemma extensional_funcset_extend_domain_inj_onI: assumes "x \ S" shows "inj_on (\(y, g). g(x := y)) {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" using assms apply (auto intro!: inj_onI) apply (metis fun_upd_same) apply (metis assms PiE_arb fun_upd_triv fun_upd_upd) done subsubsection \Misc properties of functions, composition and restriction from HOL Light\ lemma function_factors_left_gen: "(\x y. P x \ P y \ g x = g y \ f x = f y) \ (\h. \x. P x \ f x = h(g x))" (is "?lhs = ?rhs") proof assume L: ?lhs then show ?rhs apply (rule_tac x="f \ inv_into (Collect P) g" in exI) unfolding o_def by (metis (mono_tags, hide_lams) f_inv_into_f imageI inv_into_into mem_Collect_eq) qed auto lemma function_factors_left: "(\x y. (g x = g y) \ (f x = f y)) \ (\h. f = h \ g)" using function_factors_left_gen [of "\x. True" g f] unfolding o_def by blast lemma function_factors_right_gen: "(\x. P x \ (\y. g y = f x)) \ (\h. \x. P x \ f x = g(h x))" by metis lemma function_factors_right: "(\x. \y. g y = f x) \ (\h. f = g \ h)" unfolding o_def by metis lemma restrict_compose_right: "restrict (g \ restrict f S) S = restrict (g \ f) S" by auto lemma restrict_compose_left: "f ` S \ T \ restrict (restrict g T \ f) S = restrict (g \ f) S" by fastforce subsubsection \Cardinality\ lemma finite_PiE: "finite S \ (\i. i \ S \ finite (T i)) \ finite (\\<^sub>E i \ S. T i)" by (induct S arbitrary: T rule: finite_induct) (simp_all add: PiE_insert_eq) lemma inj_combinator: "x \ S \ inj_on (\(y, g). g(x := y)) (T x \ Pi\<^sub>E S T)" proof (safe intro!: inj_onI ext) fix f y g z assume "x \ S" assume fg: "f \ Pi\<^sub>E S T" "g \ Pi\<^sub>E S T" assume "f(x := y) = g(x := z)" then have *: "\i. (f(x := y)) i = (g(x := z)) i" unfolding fun_eq_iff by auto from this[of x] show "y = z" by simp fix i from *[of i] \x \ S\ fg show "f i = g i" by (auto split: if_split_asm simp: PiE_def extensional_def) qed lemma card_PiE: "finite S \ card (\\<^sub>E i \ S. T i) = (\ i\S. card (T i))" proof (induct rule: finite_induct) case empty then show ?case by auto next case (insert x S) then show ?case by (simp add: PiE_insert_eq inj_combinator card_image card_cartesian_product) qed subsection \The pigeonhole principle\ text \ An alternative formulation of this is that for a function mapping a finite set \A\ of cardinality \m\ to a finite set \B\ of cardinality \n\, there exists an element \y \ B\ that is hit at least $\lceil \frac{m}{n}\rceil$ times. However, since we do not have real numbers or rounding yet, we state it in the following equivalent form: \ lemma pigeonhole_card: assumes "f \ A \ B" "finite A" "finite B" "B \ {}" shows "\y\B. card (f -` {y} \ A) * card B \ card A" proof - from assms have "card B > 0" by auto define M where "M = Max ((\y. card (f -` {y} \ A)) ` B)" have "A = (\y\B. f -` {y} \ A)" using assms by auto also have "card \ = (\i\B. card (f -` {i} \ A))" using assms by (subst card_UN_disjoint) auto also have "\ \ (\i\B. M)" unfolding M_def using assms by (intro sum_mono Max.coboundedI) auto also have "\ = card B * M" by simp finally have "M * card B \ card A" by (simp add: mult_ac) moreover have "M \ (\y. card (f -` {y} \ A)) ` B" unfolding M_def using assms \B \ {}\ by (intro Max_in) auto ultimately show ?thesis by blast qed end