# HG changeset patch # User paulson # Date 1532790396 -3600 # Node ID 2976a4a3b1267105ffd9b9805b7a3fa9ad67bc20 # Parent 4b367da119eda44906102b558a55ca5102aedaf8 de-applying and simplification diff -r 4b367da119ed -r 2976a4a3b126 src/HOL/Algebra/Bij.thy --- a/src/HOL/Algebra/Bij.thy Wed Jul 25 22:33:04 2018 +0200 +++ b/src/HOL/Algebra/Bij.thy Sat Jul 28 16:06:36 2018 +0100 @@ -46,15 +46,11 @@ by (simp add: Bij_def compose_inv_into_id) theorem group_BijGroup: "group (BijGroup S)" -apply (simp add: BijGroup_def) -apply (rule groupI) - apply (simp add: compose_Bij) - apply (simp add: id_Bij) - apply (simp add: compose_Bij) - apply (blast intro: compose_assoc [symmetric] dest: Bij_imp_funcset) - apply (simp add: id_Bij Bij_imp_funcset Bij_imp_extensional, simp) -apply (blast intro: Bij_compose_restrict_eq restrict_inv_into_Bij) -done + apply (simp add: BijGroup_def) + apply (rule groupI) + apply (auto simp: compose_Bij id_Bij Bij_imp_funcset Bij_imp_extensional compose_assoc [symmetric]) + apply (blast intro: Bij_compose_restrict_eq restrict_inv_into_Bij) + done subsection\Automorphisms Form a Group\ @@ -63,13 +59,18 @@ by (simp add: Bij_def bij_betw_def inv_into_into) lemma Bij_inv_into_lemma: - assumes eq: "\x y. \x \ S; y \ S\ \ h(g x y) = g (h x) (h y)" - shows "\h \ Bij S; g \ S \ S \ S; x \ S; y \ S\ - \ inv_into S h (g x y) = g (inv_into S h x) (inv_into S h y)" -apply (simp add: Bij_def bij_betw_def) -apply (subgoal_tac "\x'\S. \y'\S. x = h x' \ y = h y'", clarify) - apply (simp add: eq [symmetric] inv_f_f funcset_mem [THEN funcset_mem], blast) -done + assumes eq: "\x y. \x \ S; y \ S\ \ h(g x y) = g (h x) (h y)" + and hg: "h \ Bij S" "g \ S \ S \ S" and "x \ S" "y \ S" + shows "inv_into S h (g x y) = g (inv_into S h x) (inv_into S h y)" +proof - + have "h ` S = S" + by (metis (no_types) Bij_def Int_iff assms(2) bij_betw_def mem_Collect_eq) + with \x \ S\ \y \ S\ have "\x'\S. \y'\S. x = h x' \ y = h y'" + by auto + then show ?thesis + using assms + by (auto simp add: Bij_def bij_betw_def eq [symmetric] inv_f_f funcset_mem [THEN funcset_mem]) +qed definition @@ -94,8 +95,7 @@ lemma inv_BijGroup: "f \ Bij S \ m_inv (BijGroup S) f = (\x \ S. (inv_into S f) x)" -apply (rule group.inv_equality) -apply (rule group_BijGroup) +apply (rule group.inv_equality [OF group_BijGroup]) apply (simp_all add:BijGroup_def restrict_inv_into_Bij Bij_compose_restrict_eq) done diff -r 4b367da119ed -r 2976a4a3b126 src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Wed Jul 25 22:33:04 2018 +0200 +++ b/src/HOL/Algebra/Coset.thy Sat Jul 28 16:06:36 2018 +0100 @@ -440,45 +440,36 @@ shows "N \ G" using assms normal_inv_iff by blast -corollary (in group) normal_invE: - assumes "N \ G" - shows "subgroup N G" and "\x h. \ x \ carrier G; h \ N \ \ x \ h \ inv x \ N" - using assms normal_inv_iff apply blast - by (simp add: assms normal.inv_op_closed2) - - -lemma (in group) one_is_normal : - "{\} \ G" -proof(intro normal_invI ) +lemma (in group) one_is_normal: "{\} \ G" +proof(intro normal_invI) show "subgroup {\} G" by (simp add: subgroup_def) - show "\x h. x \ carrier G \ h \ {\} \ x \ h \ inv x \ {\}" by simp -qed +qed simp subsection\More Properties of Left Cosets\ lemma (in group) l_repr_independence: - assumes "y \ x <# H" "x \ carrier G" "subgroup H G" + assumes "y \ x <# H" "x \ carrier G" and HG: "subgroup H G" shows "x <# H = y <# H" proof - obtain h' where h': "h' \ H" "y = x \ h'" using assms(1) unfolding l_coset_def by blast hence "x \ h = y \ ((inv h') \ h)" if "h \ H" for h proof - - have f3: "h' \ carrier G" - by (meson assms(3) h'(1) subgroup.mem_carrier) - have f4: "h \ carrier G" - by (meson assms(3) subgroup.mem_carrier that) - then show ?thesis - by (metis assms(2) f3 h'(2) inv_closed inv_solve_right m_assoc m_closed) + have "h' \ carrier G" + by (meson HG h'(1) subgroup.mem_carrier) + moreover have "h \ carrier G" + by (meson HG subgroup.mem_carrier that) + ultimately show ?thesis + by (metis assms(2) h'(2) inv_closed inv_solve_right m_assoc m_closed) qed - hence "\ xh. xh \ x <# H \ xh \ y <# H" - unfolding l_coset_def by (metis (no_types, lifting) UN_iff assms(3) h'(1) subgroup_def) - moreover have "\ h. h \ H \ y \ h = x \ (h' \ h)" - using h' by (meson assms(2) assms(3) m_assoc subgroup.mem_carrier) - hence "\ yh. yh \ y <# H \ yh \ x <# H" - unfolding l_coset_def using subgroup.m_closed[OF assms(3) h'(1)] by blast + hence "\xh. xh \ x <# H \ xh \ y <# H" + unfolding l_coset_def by (metis (no_types, lifting) UN_iff HG h'(1) subgroup_def) + moreover have "\h. h \ H \ y \ h = x \ (h' \ h)" + using h' by (meson assms(2) HG m_assoc subgroup.mem_carrier) + hence "\yh. yh \ y <# H \ yh \ x <# H" + unfolding l_coset_def using subgroup.m_closed[OF HG h'(1)] by blast ultimately show ?thesis by blast qed @@ -655,8 +646,8 @@ shows "hb \ a \ (\h\H. {h \ b})" proof - interpret subgroup H G by fact - from p show ?thesis apply (rule_tac UN_I [of "hb \ ((inv ha) \ h)"]) - apply blast by (simp add: inv_solve_left m_assoc) + from p show ?thesis + by (rule_tac UN_I [of "hb \ ((inv ha) \ h)"]) (auto simp: inv_solve_left m_assoc) qed lemma (in group) rcos_disjoint: @@ -666,9 +657,8 @@ proof - interpret subgroup H G by fact from p show ?thesis - apply (simp add: RCOSETS_def r_coset_def) - apply (blast intro: rcos_equation assms sym) - done + unfolding RCOSETS_def r_coset_def + by (blast intro: rcos_equation assms sym) qed @@ -761,26 +751,26 @@ proof - interpret subgroup H G by fact show ?thesis - apply (rule equalityI) - apply (force simp add: RCOSETS_def r_coset_def) - apply (auto simp add: RCOSETS_def intro: rcos_self assms) - done + unfolding RCOSETS_def r_coset_def by auto qed lemma (in group) cosets_finite: "\c \ rcosets H; H \ carrier G; finite (carrier G)\ \ finite c" -apply (auto simp add: RCOSETS_def) -apply (simp add: r_coset_subset_G [THEN finite_subset]) -done + unfolding RCOSETS_def + by (auto simp add: r_coset_subset_G [THEN finite_subset]) text\The next two lemmas support the proof of \card_cosets_equal\.\ lemma (in group) inj_on_f: - "\H \ carrier G; a \ carrier G\ \ inj_on (\y. y \ inv a) (H #> a)" -apply (rule inj_onI) -apply (subgoal_tac "x \ carrier G \ y \ carrier G") - prefer 2 apply (blast intro: r_coset_subset_G [THEN subsetD]) -apply (simp add: subsetD) -done + assumes "H \ carrier G" and a: "a \ carrier G" + shows "inj_on (\y. y \ inv a) (H #> a)" +proof + fix x y + assume "x \ H #> a" "y \ H #> a" and xy: "x \ inv a = y \ inv a" + then have "x \ carrier G" "y \ carrier G" + using assms r_coset_subset_G by blast+ + with xy a show "x = y" + by auto +qed lemma (in group) inj_on_g: "\H \ carrier G; a \ carrier G\ \ inj_on (\y. y \ a) H" @@ -827,16 +817,17 @@ using rcosets_part_G by auto proposition (in group) lagrange_finite: - "\finite(carrier G); subgroup H G\ - \ card(rcosets H) * card(H) = order(G)" -apply (simp (no_asm_simp) add: order_def rcosets_part_G [symmetric]) -apply (subst mult.commute) -apply (rule card_partition) - apply (simp add: rcosets_subset_PowG [THEN finite_subset]) - apply (simp add: rcosets_part_G) - apply (simp add: card_rcosets_equal subgroup.subset) -apply (simp add: rcos_disjoint) -done + assumes "finite(carrier G)" and HG: "subgroup H G" + shows "card(rcosets H) * card(H) = order(G)" +proof - + have "card H * card (rcosets H) = card (\(rcosets H))" + proof (rule card_partition) + show "\c1 c2. \c1 \ rcosets H; c2 \ rcosets H; c1 \ c2\ \ c1 \ c2 = {}" + using HG rcos_disjoint by auto + qed (auto simp: assms finite_UnionD rcosets_part_G card_rcosets_equal subgroup.subset) + then show ?thesis + by (simp add: HG mult.commute order_def rcosets_part_G) +qed theorem (in group) lagrange: assumes "subgroup H G" @@ -844,29 +835,29 @@ proof (cases "finite (carrier G)") case True thus ?thesis using lagrange_finite assms by simp next - case False note inf_G = this + case False thus ?thesis proof (cases "finite H") - case False thus ?thesis using inf_G by (simp add: order_def) + case False thus ?thesis using \infinite (carrier G)\ by (simp add: order_def) next - case True note finite_H = this + case True have "infinite (rcosets H)" - proof (rule ccontr) - assume "\ infinite (rcosets H)" + proof + assume "finite (rcosets H)" hence finite_rcos: "finite (rcosets H)" by simp hence "card (\(rcosets H)) = (\R\(rcosets H). card R)" - using card_Union_disjoint[of "rcosets H"] finite_H rcos_disjoint[OF assms(1)] + using card_Union_disjoint[of "rcosets H"] \finite H\ rcos_disjoint[OF assms(1)] rcosets_finite[where ?H = H] by (simp add: assms subgroup.subset) hence "order G = (\R\(rcosets H). card R)" by (simp add: assms order_def rcosets_part_G) hence "order G = (\R\(rcosets H). card H)" using card_rcosets_equal by (simp add: assms subgroup.subset) hence "order G = (card H) * (card (rcosets H))" by simp - hence "order G \ 0" using finite_rcos finite_H assms ex_in_conv + hence "order G \ 0" using finite_rcos \finite H\ assms ex_in_conv rcosets_part_G subgroup.one_closed by fastforce - thus False using inf_G order_gt_0_iff_finite by blast + thus False using \infinite (carrier G)\ order_gt_0_iff_finite by blast qed - thus ?thesis using inf_G by (simp add: order_def) + thus ?thesis using \infinite (carrier G)\ by (simp add: order_def) qed qed @@ -908,8 +899,8 @@ theorem (in normal) factorgroup_is_group: "group (G Mod H)" -apply (simp add: FactGroup_def) -apply (rule groupI) + unfolding FactGroup_def + apply (rule groupI) apply (simp add: setmult_closed) apply (simp add: normal_imp_subgroup subgroup_in_rcosets [OF is_group]) apply (simp add: restrictI setmult_closed rcosets_assoc) @@ -922,10 +913,20 @@ by (simp add: FactGroup_def) lemma (in normal) inv_FactGroup: - "X \ carrier (G Mod H) \ inv\<^bsub>G Mod H\<^esub> X = set_inv X" -apply (rule group.inv_equality [OF factorgroup_is_group]) -apply (simp_all add: FactGroup_def setinv_closed rcosets_inv_mult_group_eq) -done + assumes "X \ carrier (G Mod H)" + shows "inv\<^bsub>G Mod H\<^esub> X = set_inv X" +proof - + have X: "X \ rcosets H" + using assms by (simp add: FactGroup_def) + moreover have "set_inv X <#> X = H" + using X by (simp add: normal.rcosets_inv_mult_group_eq normal_axioms) + moreover have "Group.group (G Mod H)" + using normal.factorgroup_is_group normal_axioms by blast + moreover have "set_inv X \ rcosets H" + by (simp add: \X \ rcosets H\ setinv_closed) + ultimately show ?thesis + by (simp add: FactGroup_def group.inv_equality) +qed text\The coset map is a homomorphism from @{term G} to the quotient group @{term "G Mod H"}\ @@ -945,15 +946,13 @@ where "kernel G H h = {x. x \ carrier G \ h x = \\<^bsub>H\<^esub>}" lemma (in group_hom) subgroup_kernel: "subgroup (kernel G H h) G" -apply (rule subgroup.intro) -apply (auto simp add: kernel_def group.intro is_group) -done + by (auto simp add: kernel_def group.intro is_group intro: subgroup.intro) text\The kernel of a homomorphism is a normal subgroup\ lemma (in group_hom) normal_kernel: "(kernel G H h) \ G" -apply (simp add: G.normal_inv_iff subgroup_kernel) -apply (simp add: kernel_def) -done + apply (simp only: G.normal_inv_iff subgroup_kernel) + apply (simp add: kernel_def) + done lemma (in group_hom) FactGroup_nonempty: assumes X: "X \ carrier (G Mod kernel G H h)" @@ -982,37 +981,40 @@ lemma (in group_hom) FactGroup_hom: "(\X. the_elem (h`X)) \ hom (G Mod (kernel G H h)) H" -apply (simp add: hom_def FactGroup_the_elem_mem normal.factorgroup_is_group [OF normal_kernel] group.axioms monoid.m_closed) -proof (intro ballI) - fix X and X' - assume X: "X \ carrier (G Mod kernel G H h)" - and X': "X' \ carrier (G Mod kernel G H h)" - then - obtain g and g' - where "g \ carrier G" and "g' \ carrier G" - and "X = kernel G H h #> g" and "X' = kernel G H h #> g'" - by (auto simp add: FactGroup_def RCOSETS_def) - hence all: "\x\X. h x = h g" "\x\X'. h x = h g'" - and Xsub: "X \ carrier G" and X'sub: "X' \ carrier G" - by (force simp add: kernel_def r_coset_def image_def)+ - hence "h ` (X <#> X') = {h g \\<^bsub>H\<^esub> h g'}" using X X' - by (auto dest!: FactGroup_nonempty intro!: image_eqI - simp add: set_mult_def - subsetD [OF Xsub] subsetD [OF X'sub]) - then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" - by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique) +proof - + have "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" + if X: "X \ carrier (G Mod kernel G H h)" and X': "X' \ carrier (G Mod kernel G H h)" for X X' + proof - + obtain g and g' + where "g \ carrier G" and "g' \ carrier G" + and "X = kernel G H h #> g" and "X' = kernel G H h #> g'" + using X X' by (auto simp add: FactGroup_def RCOSETS_def) + hence all: "\x\X. h x = h g" "\x\X'. h x = h g'" + and Xsub: "X \ carrier G" and X'sub: "X' \ carrier G" + by (force simp add: kernel_def r_coset_def image_def)+ + hence "h ` (X <#> X') = {h g \\<^bsub>H\<^esub> h g'}" using X X' + by (auto dest!: FactGroup_nonempty intro!: image_eqI + simp add: set_mult_def + subsetD [OF Xsub] subsetD [OF X'sub]) + then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" + by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique) + qed + then show ?thesis + by (simp add: hom_def FactGroup_the_elem_mem normal.factorgroup_is_group [OF normal_kernel] group.axioms monoid.m_closed) qed text\Lemma for the following injectivity result\ lemma (in group_hom) FactGroup_subset: - "\g \ carrier G; g' \ carrier G; h g = h g'\ - \ kernel G H h #> g \ kernel G H h #> g'" -apply (clarsimp simp add: kernel_def r_coset_def) -apply (rename_tac y) -apply (rule_tac x="y \ g \ inv g'" in exI) -apply (simp add: G.m_assoc) -done + assumes "g \ carrier G" "g' \ carrier G" "h g = h g'" + shows "kernel G H h #> g \ kernel G H h #> g'" + unfolding kernel_def r_coset_def +proof clarsimp + fix y + assume "y \ carrier G" "h y = \\<^bsub>H\<^esub>" + with assms show "\x. x \ carrier G \ h x = \\<^bsub>H\<^esub> \ y \ g = x \ g'" + by (rule_tac x="y \ g \ inv g'" in exI) (auto simp: G.m_assoc) +qed lemma (in group_hom) FactGroup_inj_on: "inj_on (\X. the_elem (h ` X)) (carrier (G Mod kernel G H h))" @@ -1113,13 +1115,13 @@ from hHN obtain h1 h2 where h1h2 : "h1 \ H" "h2 \ N" "h = (h1,h2)" unfolding DirProd_def by fastforce hence h1h2GK : "h1 \ carrier G" "h2 \ carrier K" - using normal_imp_subgroup subgroup.subset assms apply blast+. + using normal_imp_subgroup subgroup.subset assms by blast+ have "inv\<^bsub>G \\ K\<^esub> x = (inv\<^bsub>G\<^esub> x1,inv\<^bsub>K\<^esub> x2)" using inv_DirProd[OF group_axioms assms(1) x1x2(1)x1x2(2)] x1x2 by auto hence "x \\<^bsub>G \\ K\<^esub> h \\<^bsub>G \\ K\<^esub> inv\<^bsub>G \\ K\<^esub> x = (x1 \ h1 \ inv x1,x2 \\<^bsub>K\<^esub> h2 \\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2)" using h1h2 x1x2 h1h2GK by auto moreover have "x1 \ h1 \ inv x1 \ H" "x2 \\<^bsub>K\<^esub> h2 \\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2 \ N" - using normal_invE group.normal_invE[OF assms(1)] assms x1x2 h1h2 apply auto. + using assms x1x2 h1h2 assms by (simp_all add: normal.inv_op_closed2) hence "(x1 \ h1 \ inv x1, x2 \\<^bsub>K\<^esub> h2 \\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2)\ H \ N" by auto ultimately show " x \\<^bsub>G \\ K\<^esub> h \\<^bsub>G \\ K\<^esub> inv\<^bsub>G \\ K\<^esub> x \ H \ N" by auto qed @@ -1133,7 +1135,7 @@ proof- have R :"(\(X, Y). X \ Y) \ carrier (G Mod H) \ carrier (K Mod N) \ carrier (G \\ K Mod H \ N)" - unfolding r_coset_def Sigma_def DirProd_def FactGroup_def RCOSETS_def apply simp by blast + unfolding r_coset_def Sigma_def DirProd_def FactGroup_def RCOSETS_def by force moreover have "(\x\carrier (G Mod H). \y\carrier (K Mod N). \xa\carrier (G Mod H). \ya\carrier (K Mod N). (x <#> xa) \ (y <#>\<^bsub>K\<^esub> ya) = x \ y <#>\<^bsub>G \\ K\<^esub> xa \ ya)" unfolding set_mult_def by force @@ -1143,8 +1145,14 @@ by (metis assms(2) assms(3) normal_def partial_object.select_convs(1)) moreover have "(\(X, Y). X \ Y) ` (carrier (G Mod H) \ carrier (K Mod N)) = carrier (G \\ K Mod H \ N)" - unfolding image_def apply auto using R apply force - unfolding DirProd_def FactGroup_def RCOSETS_def r_coset_def by force + proof - + have 1: "\x a b. \a \ carrier (G Mod H); b \ carrier (K Mod N)\ \ a \ b \ carrier (G \\ K Mod H \ N)" + using R by force + have 2: "\z. z \ carrier (G \\ K Mod H \ N) \ \x\carrier (G Mod H). \y\carrier (K Mod N). z = x \ y" + unfolding DirProd_def FactGroup_def RCOSETS_def r_coset_def by force + show ?thesis + unfolding image_def by (auto simp: intro: 1 2) + qed ultimately show ?thesis unfolding iso_def hom_def bij_betw_def inj_on_def by simp qed @@ -1262,7 +1270,7 @@ have hH:"h\carrier(GH)" using hHK HK_def GH_def by auto have "(\x\carrier (GH). \h\H1. x \\<^bsub>GH\<^esub> h \\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \ H1)" - using assms normal_invE GH_def normal.inv_op_closed2 by fastforce + using assms GH_def normal.inv_op_closed2 by fastforce hence INCL_1 : "x \\<^bsub>GH\<^esub> h \\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \ H1" using xH H1K_def p2 by blast have " x \\<^bsub>GH\<^esub> h \\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \ HK" @@ -1275,7 +1283,6 @@ qed qed - lemma (in group) normal_inter_subgroup: assumes "subgroup H G" and "N \ G" @@ -1288,8 +1295,8 @@ ultimately have "N \ H \ G\carrier := K \ H\" using normal_inter[of K H N] assms(1) by blast moreover have "K \ H = H" using K_def assms subgroup.subset by blast - ultimately show "normal (N\H) (G\carrier := H\)" by auto + ultimately show "normal (N\H) (G\carrier := H\)" + by auto qed - end diff -r 4b367da119ed -r 2976a4a3b126 src/HOL/Algebra/Embedded_Algebras.thy --- a/src/HOL/Algebra/Embedded_Algebras.thy Wed Jul 25 22:33:04 2018 +0200 +++ b/src/HOL/Algebra/Embedded_Algebras.thy Sat Jul 28 16:06:36 2018 +0100 @@ -187,7 +187,7 @@ corollary Span_is_add_subgroup: "set Us \ carrier R \ subgroup (Span K Us) (add_monoid R)" - using line_extension_is_subgroup add.normal_invE(1)[OF add.one_is_normal] by (induct Us) (auto) + using line_extension_is_subgroup normal_imp_subgroup[OF add.one_is_normal] by (induct Us) (auto) lemma line_extension_smult_closed: assumes "\k v. \ k \ K; v \ E \ \ k \ v \ E" and "E \ carrier R" "a \ carrier R" @@ -246,7 +246,7 @@ lemma line_extension_of_combine_set: assumes "u \ carrier R" - shows "line_extension K u { combine Ks Us | Ks. set Ks \ K } = + shows "line_extension K u { combine Ks Us | Ks. set Ks \ K } = { combine Ks (u # Us) | Ks. set Ks \ K }" (is "?line_extension = ?combinations") proof @@ -292,7 +292,7 @@ lemma line_extension_of_combine_set_length_version: assumes "u \ carrier R" - shows "line_extension K u { combine Ks Us | Ks. length Ks = length Us \ set Ks \ K } = + shows "line_extension K u { combine Ks Us | Ks. length Ks = length Us \ set Ks \ K } = { combine Ks (u # Us) | Ks. length Ks = length (u # Us) \ set Ks \ K }" (is "?line_extension = ?combinations") proof @@ -329,16 +329,16 @@ assumes "set Us \ carrier R" and "a \ carrier R" shows "a \ Span K Us \ (\k \ K - { \ }. \Ks. set Ks \ K \ combine (k # Ks) (a # Us) = \)" (is "?in_Span \ ?exists_combine") -proof +proof assume "?in_Span" then obtain Ks where Ks: "set Ks \ K" "a = combine Ks Us" using Span_eq_combine_set[OF assms(1)] by auto hence "((\ \) \ a) \ a = combine ((\ \) # Ks) (a # Us)" by auto moreover have "((\ \) \ a) \ a = \" - using assms(2) l_minus l_neg by auto + using assms(2) l_minus l_neg by auto moreover have "\ \ \ \" - using subfieldE(6)[OF K] l_neg by force + using subfieldE(6)[OF K] l_neg by force ultimately show "?exists_combine" using subring_props(3,5) Ks(1) by (force simp del: combine.simps) next @@ -391,14 +391,14 @@ proof (induct Ks Us rule: combine.induct) case (1 k Ks u Us) hence "k \ K" and "u \ set (u # Us)" by auto - hence "k \ u \ E" + hence "k \ u \ E" using 1(4) unfolding set_mult_def by auto moreover have "K <#> set Us \ E" using 1(4) unfolding set_mult_def by auto hence "combine Ks Us \ E" using 1 by auto ultimately show ?case - using add.subgroupE(4)[OF assms(2)] by auto + using add.subgroupE(4)[OF assms(2)] by auto next case "2_1" thus ?case using subgroup.one_closed[OF assms(2)] by auto @@ -436,7 +436,7 @@ hence "combine [ k ] (u # Us) \ Span K (u # Us)" using Span_eq_combine_set[OF Cons(2)] by (auto simp del: combine.simps) moreover have "k \ carrier R" and "u \ carrier R" - using Cons(2) k subring_props(1) by (blast, auto) + using Cons(2) k subring_props(1) by (blast, auto) ultimately show "k \ u \ Span K (u # Us)" by (auto simp del: Span.simps) qed @@ -455,7 +455,7 @@ corollary Span_same_set: assumes "set Us \ carrier R" shows "set Us = set Vs \ Span K Us = Span K Vs" - using Span_eq_generate assms by auto + using Span_eq_generate assms by auto corollary Span_incl: "set Us \ carrier R \ K <#> (set Us) \ Span K Us" using Span_eq_generate generate.incl[of _ _ "add_monoid R"] by auto @@ -583,7 +583,7 @@ moreover have "Span K Us \ Span K (u # Us)" using mono_Span independent_in_carrier[OF assms] by auto ultimately show ?thesis - using independent_backwards(1)[OF assms] by auto + using independent_backwards(1)[OF assms] by auto qed corollary independent_replacement: @@ -624,7 +624,7 @@ from assms show "Span K Us \ Span K Vs = { \ }" proof (induct Us rule: list.induct) case Nil thus ?case - using Span_subgroup_props(2)[OF independent_in_carrier[of K Vs]] by simp + using Span_subgroup_props(2)[OF independent_in_carrier[of K Vs]] by simp next case (Cons u Us) hence IH: "Span K Us \ Span K Vs = {\}" by auto @@ -653,7 +653,7 @@ hence "k \ u = (\ u') \ v'" using in_carrier(1) k(2) u'(2) v'(2) add.m_comm r_neg1 by auto hence "k \ u \ Span K (Us @ Vs)" - using Span_subgroup_props(4)[OF in_carrier(2) u'(1)] v'(1) + using Span_subgroup_props(4)[OF in_carrier(2) u'(1)] v'(1) Span_append_eq_set_add[OF in_carrier(2-3)] unfolding set_add_def' by blast hence "u \ Span K (Us @ Vs)" using Cons(2) Span_m_inv_simprule[OF _ _ in_carrier(1), of "Us @ Vs" k] @@ -678,7 +678,7 @@ hence in_carrier: "u \ carrier R" "set Us \ carrier R" "set Vs \ carrier R" "set (u # Us) \ carrier R" using Cons(2-3)[THEN independent_in_carrier] by auto - hence "Span K Us \ Span K (u # Us)" + hence "Span K Us \ Span K (u # Us)" using mono_Span by auto hence "Span K Us \ Span K Vs = { \ }" using Cons(4) Span_subgroup_props(2)[OF in_carrier(2)] by auto @@ -733,7 +733,7 @@ hence "combine Ks' Us = \" using combine_in_carrier[OF _ Us, of Ks'] Ks' u Cons(3) subring_props(1) unfolding Ks by auto hence "set (take (length Us) Ks') \ { \ }" - using Cons(1)[OF Ks' _ independent_backwards(2)[OF Cons(4)]] by simp + using Cons(1)[OF Ks' _ independent_backwards(2)[OF Cons(4)]] by simp thus ?thesis using k_zero unfolding Ks by auto qed @@ -878,10 +878,10 @@ case (Cons u Us) then obtain Vs' Vs'' where Vs: "Vs = Vs' @ (u # Vs'')" by (metis list.set_intros(1) split_list) - + have in_carrier: "u \ carrier R" "set Us \ carrier R" - using independent_in_carrier[OF Cons(2)] by auto - + using independent_in_carrier[OF Cons(2)] by auto + have "distinct Vs" using Cons(3-4) independent_distinct[OF Cons(2)] by (metis card_distinct distinct_card) @@ -905,7 +905,7 @@ shows "\Vs'. set Vs' \ set Vs \ length Vs' = length Us' \ independent K (Vs' @ Us)" using assms proof (induct "length Us'" arbitrary: Us' Us) - case 0 thus ?case by auto + case 0 thus ?case by auto next case (Suc n) then obtain u Us'' where Us'': "Us' = Us'' @ [u]" @@ -1074,9 +1074,9 @@ using space_subgroup_props(1)[OF assms(1)] li_Cons[OF _ v(2) step(4)] by auto then obtain Vs where "length (Vs @ (v # Us)) = n" "independent K (Vs @ (v # Us))" "Span K (Vs @ (v # Us)) = E" - using step(3)[of "v # Us"] step(1-2,4-6) v by auto + using step(3)[of "v # Us"] step(1-2,4-6) v by auto thus ?case - by (metis append.assoc append_Cons append_Nil) + by (metis append.assoc append_Cons append_Nil) qed } note aux_lemma = this have "length Us \ n" @@ -1119,7 +1119,7 @@ hence in_carrier: "set Us \ carrier R" "set (Vs @ Bs) \ carrier R" using independent_in_carrier[OF Us(2)] independent_in_carrier[OF Vs(2)] by auto hence "Span K Us \ (Span K (Vs @ Bs)) \ Span K Bs" - using Bs(4) Us(3) Vs(3) mono_Span_append(1)[OF _ Bs(1), of Us] by auto + using Bs(4) Us(3) Vs(3) mono_Span_append(1)[OF _ Bs(1), of Us] by auto hence "Span K Us \ (Span K (Vs @ Bs)) \ { \ }" using independent_split(3)[OF Us(2)] by blast hence "Span K Us \ (Span K (Vs @ Bs)) = { \ }" @@ -1147,7 +1147,7 @@ ultimately show "v \ (Span K Us) <+>\<^bsub>R\<^esub> F" using u1' unfolding set_add_def' by auto qed - ultimately have "Span K (Us @ (Vs @ Bs)) = E <+>\<^bsub>R\<^esub> F" + ultimately have "Span K (Us @ (Vs @ Bs)) = E <+>\<^bsub>R\<^esub> F" using Span_append_eq_set_add[OF in_carrier] Vs(3) by auto thus ?thesis using dim by simp @@ -1169,7 +1169,7 @@ by (metis One_nat_def length_0_conv length_Suc_conv) have in_carrier: "set (map (\u'. u' \ u) Us) \ carrier R" using Us(1) u(1) by (induct Us) (auto) - + have li: "independent K (map (\u'. u' \ u) Us)" proof (rule trivial_combine_imp_independent[OF assms(1) in_carrier]) fix Ks assume Ks: "set Ks \ K" and "combine Ks (map (\u'. u' \ u) Us) = \" @@ -1244,7 +1244,7 @@ ultimately have "dimension (n * Suc m) K (Span F [ v ] <+>\<^bsub>R\<^esub> Span F Vs')" using dimension_direct_sum_space[OF assms(1) _ _ inter] by auto thus "dimension (n * Suc m) K E" - using Span_append_eq_set_add[OF assms(2) li[THEN independent_in_carrier]] Vs(4) v by auto + using Span_append_eq_set_add[OF assms(2) li[THEN independent_in_carrier]] Vs(4) v by auto qed @@ -1271,14 +1271,14 @@ hence "combine Ks Us = (combine (take (length Us) Ks) Us) \ \" using combine_append[OF _ _ assms(2), of "take (length Us) Ks" "drop (length Us) Ks" "[]"] len by auto also have " ... = combine (take (length Us) Ks) Us" - using combine_in_carrier[OF set_t assms(2)] by auto + using combine_in_carrier[OF set_t assms(2)] by auto finally show "combine Ks Us = combine (take (length Us) Ks) Us" . qed *) (* lemma combine_normalize: - assumes "set Ks \ K" "set Us \ carrier R" "a = combine Ks Us" + assumes "set Ks \ K" "set Us \ carrier R" "a = combine Ks Us" shows "\Ks'. set Ks' \ K \ length Ks' = length Us \ a = combine Ks' Us" proof (cases "length Ks \ length Us") assume "\ length Ks \ length Us" @@ -1291,12 +1291,12 @@ next assume len: "length Ks \ length Us" have Ks: "set Ks \ carrier R" and set_r: "set (replicate (length Us - length Ks) \) \ carrier R" - using assms subring_props(1) zero_closed by (metis dual_order.trans, auto) + using assms subring_props(1) zero_closed by (metis dual_order.trans, auto) moreover have set_t: "set (take (length Ks) Us) \ carrier R" and set_d: "set (drop (length Ks) Us) \ carrier R" using assms(2) len dual_order.trans by (metis set_take_subset, metis set_drop_subset) - ultimately + ultimately have "combine (Ks @ (replicate (length Us - length Ks) \)) Us = (combine Ks (take (length Ks) Us)) \ (combine (replicate (length Us - length Ks) \) (drop (length Ks) Us))" diff -r 4b367da119ed -r 2976a4a3b126 src/HOL/Algebra/Generated_Groups.thy --- a/src/HOL/Algebra/Generated_Groups.thy Wed Jul 25 22:33:04 2018 +0200 +++ b/src/HOL/Algebra/Generated_Groups.thy Sat Jul 28 16:06:36 2018 +0100 @@ -466,7 +466,7 @@ shows "derived G H \ G" unfolding derived_def proof (rule normal_generateI) show "(\h1 \ H. \h2 \ H. { h1 \ h2 \ inv h1 \ inv h2 }) \ carrier G" - using subgroup.subset assms normal_invE(1) by blast + using subgroup.subset assms normal_imp_subgroup by blast next show "\h g. \ h \ derived_set G H; g \ carrier G \ \ g \ h \ inv g \ derived_set G H" proof - @@ -474,7 +474,7 @@ then obtain h1 h2 where h1: "h1 \ H" "h1 \ carrier G" and h2: "h2 \ H" "h2 \ carrier G" and h: "h = h1 \ h2 \ inv h1 \ inv h2" - using subgroup.subset assms normal_invE(1) by blast + using subgroup.subset assms normal_imp_subgroup by blast hence "g \ h \ inv g = g \ h1 \ (inv g \ g) \ h2 \ (inv g \ g) \ inv h1 \ (inv g \ g) \ inv h2 \ inv g" by (simp add: g m_assoc) @@ -486,8 +486,8 @@ have "g \ h \ inv g = (g \ h1 \ inv g) \ (g \ h2 \ inv g) \ inv (g \ h1 \ inv g) \ inv (g \ h2 \ inv g)" by (simp add: g h1 h2 inv_mult_group m_assoc) - moreover have "g \ h1 \ inv g \ H" by (simp add: assms normal_invE(2) g h1) - moreover have "g \ h2 \ inv g \ H" by (simp add: assms normal_invE(2) g h2) + moreover have "g \ h1 \ inv g \ H" by (simp add: assms normal.inv_op_closed2 g h1) + moreover have "g \ h2 \ inv g \ H" by (simp add: assms normal.inv_op_closed2 g h2) ultimately show "g \ h \ inv g \ derived_set G H" by blast qed qed diff -r 4b367da119ed -r 2976a4a3b126 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Wed Jul 25 22:33:04 2018 +0200 +++ b/src/HOL/Algebra/Group.thy Sat Jul 28 16:06:36 2018 +0100 @@ -763,13 +763,13 @@ and "subgroup I K" shows "subgroup (H \ I) (G \\ K)" proof (intro group.group_incl_imp_subgroup[OF DirProd_group[OF assms(1)assms(3)]]) - have "H \ carrier G" "I \ carrier K" using subgroup.subset assms apply blast+. + have "H \ carrier G" "I \ carrier K" using subgroup.subset assms by blast+ thus "(H \ I) \ carrier (G \\ K)" unfolding DirProd_def by auto have "Group.group ((G\carrier := H\) \\ (K\carrier := I\))" using DirProd_group[OF subgroup.subgroup_is_group[OF assms(2)assms(1)] subgroup.subgroup_is_group[OF assms(4)assms(3)]]. moreover have "((G\carrier := H\) \\ (K\carrier := I\)) = ((G \\ K)\carrier := H \ I\)" - unfolding DirProd_def using assms apply simp. + unfolding DirProd_def using assms by simp ultimately show "Group.group ((G \\ K)\carrier := H \ I\)" by simp qed @@ -1054,7 +1054,7 @@ lemma (in group_hom) img_is_subgroup: "subgroup (h ` (carrier G)) H" apply (rule subgroupI) apply (auto simp add: image_subsetI) - apply (metis (no_types, hide_lams) G.inv_closed hom_inv image_iff) + apply (metis G.inv_closed hom_inv image_iff) by (metis G.monoid_axioms hom_mult image_eqI monoid.m_closed) lemma (in group_hom) subgroup_img_is_subgroup: @@ -1157,9 +1157,8 @@ show "monoid (G\carrier := H\)" using submonoid.submonoid_is_monoid assms comm_monoid_axioms comm_monoid_def by blast show "\x y. x \ carrier (G\carrier := H\) \ y \ carrier (G\carrier := H\) - \ x \\<^bsub>G\carrier := H\\<^esub> y = y \\<^bsub>G\carrier := H\\<^esub> x" apply simp - using assms comm_monoid_axioms_def submonoid.mem_carrier - by (metis m_comm) + \ x \\<^bsub>G\carrier := H\\<^esub> y = y \\<^bsub>G\carrier := H\\<^esub> x" + by simp (meson assms m_comm submonoid.mem_carrier) qed locale comm_group = comm_monoid + group diff -r 4b367da119ed -r 2976a4a3b126 src/HOL/Algebra/Ideal.thy --- a/src/HOL/Algebra/Ideal.thy Wed Jul 25 22:33:04 2018 +0200 +++ b/src/HOL/Algebra/Ideal.thy Sat Jul 28 16:06:36 2018 +0100 @@ -128,10 +128,9 @@ proof - interpret additive_subgroup I R by fact interpret cring R by fact - show ?thesis apply intro_locales - apply (intro ideal_axioms.intro) - apply (erule (1) I_l_closed) - apply (erule (1) I_r_closed) + show ?thesis + apply intro_locales + apply (simp add: I_l_closed I_r_closed ideal_axioms_def) by (simp add: I_notcarr I_prime primeideal_axioms.intro) qed diff -r 4b367da119ed -r 2976a4a3b126 src/HOL/Algebra/RingHom.thy --- a/src/HOL/Algebra/RingHom.thy Wed Jul 25 22:33:04 2018 +0200 +++ b/src/HOL/Algebra/RingHom.thy Sat Jul 28 16:06:36 2018 +0100 @@ -20,11 +20,10 @@ by standard (rule homh) sublocale ring_hom_ring \ abelian_group?: abelian_group_hom R S -apply (intro abelian_group_homI R.is_abelian_group S.is_abelian_group) -apply (intro group_hom.intro group_hom_axioms.intro R.a_group S.a_group) -apply (insert homh, unfold hom_def ring_hom_def) -apply simp -done +proof + show "h \ hom (add_monoid R) (add_monoid S)" + using homh by (simp add: hom_def ring_hom_def) +qed lemma (in ring_hom_ring) is_ring_hom_ring: "ring_hom_ring R S h" @@ -33,8 +32,7 @@ lemma ring_hom_ringI: fixes R (structure) and S (structure) assumes "ring R" "ring S" - assumes (* morphism: "h \ carrier R \ carrier S" *) - hom_closed: "!!x. x \ carrier R ==> h x \ carrier S" + assumes hom_closed: "!!x. x \ carrier R ==> h x \ carrier S" and compatible_mult: "\x y. [| x \ carrier R; y \ carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_add: "\x y. [| x \ carrier R; y \ carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_one: "h \ = \\<^bsub>S\<^esub>" @@ -42,13 +40,12 @@ proof - interpret ring R by fact interpret ring S by fact - show ?thesis apply unfold_locales -apply (unfold ring_hom_def, safe) - apply (simp add: hom_closed Pi_def) - apply (erule (1) compatible_mult) - apply (erule (1) compatible_add) -apply (rule compatible_one) -done + show ?thesis + proof + show "h \ ring_hom R S" + unfolding ring_hom_def + by (auto simp: compatible_mult compatible_add compatible_one hom_closed) + qed qed lemma ring_hom_ringI2: @@ -58,11 +55,11 @@ proof - interpret R: ring R by fact interpret S: ring S by fact - show ?thesis apply (intro ring_hom_ring.intro ring_hom_ring_axioms.intro) - apply (rule R.is_ring) - apply (rule S.is_ring) - apply (rule h) - done + show ?thesis + proof + show "h \ ring_hom R S" + using h . + qed qed lemma ring_hom_ringI3: @@ -75,13 +72,11 @@ interpret abelian_group_hom R S h by fact interpret R: ring R by fact interpret S: ring S by fact - show ?thesis apply (intro ring_hom_ring.intro ring_hom_ring_axioms.intro, rule R.is_ring, rule S.is_ring) - apply (insert group_hom.homh[OF a_group_hom]) - apply (unfold hom_def ring_hom_def, simp) - apply safe - apply (erule (1) compatible_mult) - apply (rule compatible_one) - done + show ?thesis + proof + show "h \ ring_hom R S" + unfolding ring_hom_def by (auto simp: compatible_one compatible_mult) + qed qed lemma ring_hom_cringI: @@ -91,21 +86,22 @@ interpret ring_hom_ring R S h by fact interpret R: cring R by fact interpret S: cring S by fact - show ?thesis by (intro ring_hom_cring.intro ring_hom_cring_axioms.intro) - (rule R.is_cring, rule S.is_cring, rule homh) + show ?thesis + proof + show "h \ ring_hom R S" + by (simp add: homh) + qed qed subsection \The Kernel of a Ring Homomorphism\ \ \the kernel of a ring homomorphism is an ideal\ -lemma (in ring_hom_ring) kernel_is_ideal: - shows "ideal (a_kernel R S h) R" -apply (rule idealI) - apply (rule R.is_ring) - apply (rule additive_subgroup.a_subgroup[OF additive_subgroup_a_kernel]) - apply (unfold a_kernel_def', simp+) -done +lemma (in ring_hom_ring) kernel_is_ideal: "ideal (a_kernel R S h) R" + apply (rule idealI [OF R.is_ring]) + apply (rule additive_subgroup.a_subgroup[OF additive_subgroup_a_kernel]) + apply (auto simp: a_kernel_def') + done text \Elements of the kernel are mapped to zero\ lemma (in abelian_group_hom) kernel_zero [simp]: @@ -174,29 +170,10 @@ corollary (in ring_hom_ring) rcos_eq_homeq: assumes acarr: "a \ carrier R" shows "(a_kernel R S h) +> a = {x \ carrier R. h x = h a}" - apply rule defer 1 - apply clarsimp defer 1 -proof +proof - interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal) - - fix x - assume xrcos: "x \ a_kernel R S h +> a" - from acarr and this - have xcarr: "x \ carrier R" - by (rule a_elemrcos_carrier) - - from xrcos - have "h x = h a" by (rule rcos_imp_homeq[OF acarr]) - from xcarr and this - show "x \ {x \ carrier R. h x = h a}" by fast -next - interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal) - - fix x - assume xcarr: "x \ carrier R" - and hx: "h x = h a" - from acarr xcarr hx - show "x \ a_kernel R S h +> a" by (rule homeq_imp_rcos) + show ?thesis + using assms by (auto simp: intro: homeq_imp_rcos rcos_imp_homeq a_elemrcos_carrier) qed lemma (in ring_hom_ring) nat_pow_hom: diff -r 4b367da119ed -r 2976a4a3b126 src/HOL/Library/FuncSet.thy --- a/src/HOL/Library/FuncSet.thy Wed Jul 25 22:33:04 2018 +0200 +++ b/src/HOL/Library/FuncSet.thy Sat Jul 28 16:06:36 2018 +0100 @@ -163,8 +163,6 @@ lemma compose_assoc: assumes "f \ A \ B" - and "g \ B \ C" - and "h \ C \ D" 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)