| author | wenzelm | 
| Sat, 08 Sep 2018 22:43:25 +0200 | |
| changeset 68955 | 0851db8cde12 | 
| parent 68687 | 2976a4a3b126 | 
| child 68975 | 5ce4d117cea7 | 
| permissions | -rw-r--r-- | 
| 14706 | 1 | (* Title: HOL/Algebra/Coset.thy | 
| 68582 | 2 | Authors: Florian Kammueller, L C Paulson, Stephan Hohe | 
| 3 | ||
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 4 | With additional contributions from Martin Baillon and Paulo Emílio de Vilhena. | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 5 | *) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 6 | |
| 35849 | 7 | theory Coset | 
| 8 | imports Group | |
| 9 | begin | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 10 | |
| 61382 | 11 | section \<open>Cosets and Quotient Groups\<close> | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 12 | |
| 35847 | 13 | definition | 
| 14963 | 14 | r_coset :: "[_, 'a set, 'a] \<Rightarrow> 'a set" (infixl "#>\<index>" 60) | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 15 |   where "H #>\<^bsub>G\<^esub> a = (\<Union>h\<in>H. {h \<otimes>\<^bsub>G\<^esub> a})"
 | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 16 | |
| 35847 | 17 | definition | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 18 | l_coset :: "[_, 'a, 'a set] \<Rightarrow> 'a set" (infixl "<#\<index>" 60) | 
| 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 19 |   where "a <#\<^bsub>G\<^esub> H = (\<Union>h\<in>H. {a \<otimes>\<^bsub>G\<^esub> h})"
 | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 20 | |
| 35847 | 21 | definition | 
| 14963 | 22 |   RCOSETS  :: "[_, 'a set] \<Rightarrow> ('a set)set"   ("rcosets\<index> _" [81] 80)
 | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 23 |   where "rcosets\<^bsub>G\<^esub> H = (\<Union>a\<in>carrier G. {H #>\<^bsub>G\<^esub> a})"
 | 
| 14963 | 24 | |
| 35847 | 25 | definition | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 26 | set_mult :: "[_, 'a set ,'a set] \<Rightarrow> 'a set" (infixl "<#>\<index>" 60) | 
| 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 27 |   where "H <#>\<^bsub>G\<^esub> K = (\<Union>h\<in>H. \<Union>k\<in>K. {h \<otimes>\<^bsub>G\<^esub> k})"
 | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 28 | |
| 35847 | 29 | definition | 
| 14963 | 30 |   SET_INV :: "[_,'a set] \<Rightarrow> 'a set"  ("set'_inv\<index> _" [81] 80)
 | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 31 |   where "set_inv\<^bsub>G\<^esub> H = (\<Union>h\<in>H. {inv\<^bsub>G\<^esub> h})"
 | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 32 | |
| 14963 | 33 | |
| 34 | locale normal = subgroup + group + | |
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 35 | assumes coset_eq: "(\<forall>x \<in> carrier G. H #> x = x <# H)" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 36 | |
| 19380 | 37 | abbreviation | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20318diff
changeset | 38 |   normal_rel :: "['a set, ('a, 'b) monoid_scheme] \<Rightarrow> bool"  (infixl "\<lhd>" 60) where
 | 
| 19380 | 39 | "H \<lhd> G \<equiv> normal H G" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 40 | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 41 | (*Next two lemmas contributed by Martin Baillon.*) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 42 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 43 | lemma l_coset_eq_set_mult: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 44 | fixes G (structure) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 45 |   shows "x <# H = {x} <#> H"
 | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 46 | unfolding l_coset_def set_mult_def by simp | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 47 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 48 | lemma r_coset_eq_set_mult: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 49 | fixes G (structure) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 50 |   shows "H #> x = H <#> {x}"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 51 | unfolding r_coset_def set_mult_def by simp | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 52 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 53 | (* Next five lemmas contributed by Paulo Emílio de Vilhena. *) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 54 | |
| 68517 | 55 | lemma (in subgroup) rcosets_non_empty: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 56 | assumes "R \<in> rcosets H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 57 |   shows "R \<noteq> {}"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 58 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 59 | obtain g where "g \<in> carrier G" "R = H #> g" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 60 | using assms unfolding RCOSETS_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 61 | hence "\<one> \<otimes> g \<in> R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 62 | using one_closed unfolding r_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 63 | thus ?thesis by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 64 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 65 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 66 | lemma (in group) diff_neutralizes: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 67 | assumes "subgroup H G" "R \<in> rcosets H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 68 | shows "\<And>r1 r2. \<lbrakk> r1 \<in> R; r2 \<in> R \<rbrakk> \<Longrightarrow> r1 \<otimes> (inv r2) \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 69 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 70 | fix r1 r2 assume r1: "r1 \<in> R" and r2: "r2 \<in> R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 71 | obtain g where g: "g \<in> carrier G" "R = H #> g" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 72 | using assms unfolding RCOSETS_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 73 | then obtain h1 h2 where h1: "h1 \<in> H" "r1 = h1 \<otimes> g" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 74 | and h2: "h2 \<in> H" "r2 = h2 \<otimes> g" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 75 | using r1 r2 unfolding r_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 76 | hence "r1 \<otimes> (inv r2) = (h1 \<otimes> g) \<otimes> ((inv g) \<otimes> (inv h2))" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 77 | using inv_mult_group is_group assms(1) g(1) subgroup.mem_carrier by fastforce | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 78 | also have " ... = (h1 \<otimes> (g \<otimes> inv g) \<otimes> inv h2)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 79 | using h1 h2 assms(1) g(1) inv_closed m_closed monoid.m_assoc | 
| 68604 | 80 | monoid_axioms subgroup.mem_carrier | 
| 81 | proof - | |
| 82 | have "h1 \<in> carrier G" | |
| 83 | by (meson subgroup.mem_carrier assms(1) h1(1)) | |
| 84 | moreover have "h2 \<in> carrier G" | |
| 85 | by (meson subgroup.mem_carrier assms(1) h2(1)) | |
| 86 | ultimately show ?thesis | |
| 87 | using g(1) inv_closed m_assoc m_closed by presburger | |
| 88 | qed | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 89 | finally have "r1 \<otimes> inv r2 = h1 \<otimes> inv h2" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 90 | using assms(1) g(1) h1(1) subgroup.mem_carrier by fastforce | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 91 | thus "r1 \<otimes> inv r2 \<in> H" by (metis assms(1) h1(1) h2(1) subgroup_def) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 92 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 93 | |
| 68517 | 94 | lemma mono_set_mult: "\<lbrakk> H \<subseteq> H'; K \<subseteq> K' \<rbrakk> \<Longrightarrow> H <#>\<^bsub>G\<^esub> K \<subseteq> H' <#>\<^bsub>G\<^esub> K'" | 
| 95 | unfolding set_mult_def by (simp add: UN_mono) | |
| 96 | ||
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 97 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 98 | subsection \<open>Stable Operations for Subgroups\<close> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 99 | |
| 68517 | 100 | lemma set_mult_consistent [simp]: | 
| 101 | "N <#>\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> K = N <#>\<^bsub>G\<^esub> K" | |
| 102 | unfolding set_mult_def by simp | |
| 103 | ||
| 104 | lemma r_coset_consistent [simp]: | |
| 105 | "I #>\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> h = I #>\<^bsub>G\<^esub> h" | |
| 106 | unfolding r_coset_def by simp | |
| 107 | ||
| 108 | lemma l_coset_consistent [simp]: | |
| 109 | "h <#\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> I = h <#\<^bsub>G\<^esub> I" | |
| 110 | unfolding l_coset_def by simp | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 111 | |
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 112 | subsection \<open>Basic Properties of set multiplication\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 113 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 114 | lemma (in group) setmult_subset_G: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 115 | assumes "H \<subseteq> carrier G" "K \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 116 | shows "H <#> K \<subseteq> carrier G" using assms | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 117 | by (auto simp add: set_mult_def subsetD) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 118 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 119 | lemma (in monoid) set_mult_closed: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 120 | assumes "H \<subseteq> carrier G" "K \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 121 | shows "H <#> K \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 122 | using assms by (auto simp add: set_mult_def subsetD) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 123 | |
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 124 | (* Next lemma contributed by Martin Baillon.*) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 125 | lemma (in group) set_mult_assoc: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 126 | assumes "M \<subseteq> carrier G" "H \<subseteq> carrier G" "K \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 127 | shows "(M <#> H) <#> K = M <#> (H <#> K)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 128 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 129 | show "(M <#> H) <#> K \<subseteq> M <#> (H <#> K)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 130 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 131 | fix x assume "x \<in> (M <#> H) <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 132 | then obtain m h k where x: "m \<in> M" "h \<in> H" "k \<in> K" "x = (m \<otimes> h) \<otimes> k" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 133 | unfolding set_mult_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 134 | hence "x = m \<otimes> (h \<otimes> k)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 135 | using assms m_assoc by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 136 | thus "x \<in> M <#> (H <#> K)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 137 | unfolding set_mult_def using x by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 138 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 139 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 140 | show "M <#> (H <#> K) \<subseteq> (M <#> H) <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 141 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 142 | fix x assume "x \<in> M <#> (H <#> K)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 143 | then obtain m h k where x: "m \<in> M" "h \<in> H" "k \<in> K" "x = m \<otimes> (h \<otimes> k)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 144 | unfolding set_mult_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 145 | hence "x = (m \<otimes> h) \<otimes> k" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 146 | using assms m_assoc rev_subsetD by metis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 147 | thus "x \<in> (M <#> H) <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 148 | unfolding set_mult_def using x by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 149 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 150 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 151 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 152 | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 153 | |
| 61382 | 154 | subsection \<open>Basic Properties of Cosets\<close> | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 155 | |
| 14747 | 156 | lemma (in group) coset_mult_assoc: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 157 | assumes "M \<subseteq> carrier G" "g \<in> carrier G" "h \<in> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 158 | shows "(M #> g) #> h = M #> (g \<otimes> h)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 159 | using assms by (force simp add: r_coset_def m_assoc) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 160 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 161 | lemma (in group) coset_assoc: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 162 | assumes "x \<in> carrier G" "y \<in> carrier G" "H \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 163 | shows "x <# (H #> y) = (x <# H) #> y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 164 |   using set_mult_assoc[of "{x}" H "{y}"]
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 165 | by (simp add: l_coset_eq_set_mult r_coset_eq_set_mult assms) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 166 | |
| 14747 | 167 | lemma (in group) coset_mult_one [simp]: "M \<subseteq> carrier G ==> M #> \<one> = M" | 
| 168 | by (force simp add: r_coset_def) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 169 | |
| 14747 | 170 | lemma (in group) coset_mult_inv1: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 171 | assumes "M #> (x \<otimes> (inv y)) = M" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 172 | and "x \<in> carrier G" "y \<in> carrier G" "M \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 173 | shows "M #> x = M #> y" using assms | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 174 | by (metis coset_mult_assoc group.inv_solve_right is_group subgroup_def subgroup_self) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 175 | |
| 14747 | 176 | lemma (in group) coset_mult_inv2: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 177 | assumes "M #> x = M #> y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 178 | and "x \<in> carrier G" "y \<in> carrier G" "M \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 179 | shows "M #> (x \<otimes> (inv y)) = M " using assms | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 180 | by (metis group.coset_mult_assoc group.coset_mult_one inv_closed is_group r_inv) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 181 | |
| 14747 | 182 | lemma (in group) coset_join1: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 183 | assumes "H #> x = H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 184 | and "x \<in> carrier G" "subgroup H G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 185 | shows "x \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 186 | using assms r_coset_def l_one subgroup.one_closed sym by fastforce | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 187 | |
| 14747 | 188 | lemma (in group) solve_equation: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 189 | assumes "subgroup H G" "x \<in> H" "y \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 190 | shows "\<exists>h \<in> H. y = h \<otimes> x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 191 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 192 | have "y = (y \<otimes> (inv x)) \<otimes> x" using assms | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 193 | by (simp add: m_assoc subgroup.mem_carrier) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 194 | moreover have "y \<otimes> (inv x) \<in> H" using assms | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 195 | by (simp add: subgroup_def) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 196 | ultimately show ?thesis by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 197 | qed | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 198 | |
| 14963 | 199 | lemma (in group) repr_independence: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 200 | assumes "y \<in> H #> x" "x \<in> carrier G" "subgroup H G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 201 | shows "H #> x = H #> y" using assms | 
| 14963 | 202 | by (auto simp add: r_coset_def m_assoc [symmetric] | 
| 203 | subgroup.subset [THEN subsetD] | |
| 204 | subgroup.m_closed solve_equation) | |
| 205 | ||
| 14747 | 206 | lemma (in group) coset_join2: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 207 | assumes "x \<in> carrier G" "subgroup H G" "x \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 208 | shows "H #> x = H" using assms | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67091diff
changeset | 209 |   \<comment> \<open>Alternative proof is to put @{term "x=\<one>"} in \<open>repr_independence\<close>.\<close>
 | 
| 14963 | 210 | by (force simp add: subgroup.m_closed r_coset_def solve_equation) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 211 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 212 | lemma (in group) coset_join3: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 213 | assumes "x \<in> carrier G" "subgroup H G" "x \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 214 | shows "x <# H = H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 215 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 216 | have "\<And>h. h \<in> H \<Longrightarrow> x \<otimes> h \<in> H" using assms | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 217 | by (simp add: subgroup.m_closed) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 218 | thus "x <# H \<subseteq> H" unfolding l_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 219 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 220 | have "\<And>h. h \<in> H \<Longrightarrow> x \<otimes> ((inv x) \<otimes> h) = h" | 
| 68604 | 221 | by (metis (no_types, lifting) assms group.inv_closed group.inv_solve_left is_group | 
| 222 | monoid.m_closed monoid_axioms subgroup.mem_carrier) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 223 | moreover have "\<And>h. h \<in> H \<Longrightarrow> (inv x) \<otimes> h \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 224 | by (simp add: assms subgroup.m_closed subgroup.m_inv_closed) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 225 | ultimately show "H \<subseteq> x <# H" unfolding l_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 226 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 227 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 228 | lemma (in monoid) r_coset_subset_G: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 229 | "\<lbrakk> H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> H #> x \<subseteq> carrier G" | 
| 14747 | 230 | by (auto simp add: r_coset_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 231 | |
| 14747 | 232 | lemma (in group) rcosI: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 233 | "\<lbrakk> h \<in> H; H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> h \<otimes> x \<in> H #> x" | 
| 14747 | 234 | by (auto simp add: r_coset_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 235 | |
| 14963 | 236 | lemma (in group) rcosetsI: | 
| 237 | "\<lbrakk>H \<subseteq> carrier G; x \<in> carrier G\<rbrakk> \<Longrightarrow> H #> x \<in> rcosets H" | |
| 238 | by (auto simp add: RCOSETS_def) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 239 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 240 | lemma (in group) rcos_self: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 241 | "\<lbrakk> x \<in> carrier G; subgroup H G \<rbrakk> \<Longrightarrow> x \<in> H #> x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 242 | by (metis l_one rcosI subgroup_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 243 | |
| 61382 | 244 | text (in group) \<open>Opposite of @{thm [source] "repr_independence"}\<close>
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 245 | lemma (in group) repr_independenceD: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 246 | assumes "subgroup H G" "y \<in> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 247 | and "H #> x = H #> y" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 248 | shows "y \<in> H #> x" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 249 | using assms by (simp add: rcos_self) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 250 | |
| 61382 | 251 | text \<open>Elements of a right coset are in the carrier\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 252 | lemma (in subgroup) elemrcos_carrier: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 253 | assumes "group G" "a \<in> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 254 | and "a' \<in> H #> a" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 255 | shows "a' \<in> carrier G" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 256 | by (meson assms group.is_monoid monoid.r_coset_subset_G subset subsetCE) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 257 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 258 | lemma (in subgroup) rcos_const: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 259 | assumes "group G" "h \<in> H" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 260 | shows "H #> h = H" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 261 | using group.coset_join2[OF assms(1), of h H] | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 262 | by (simp add: assms(2) subgroup_axioms) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 263 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 264 | lemma (in subgroup) rcos_module_imp: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 265 | assumes "group G" "x \<in> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 266 | and "x' \<in> H #> x" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 267 | shows "(x' \<otimes> inv x) \<in> H" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 268 | proof - | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 269 | obtain h where h: "h \<in> H" "x' = h \<otimes> x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 270 | using assms(3) unfolding r_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 271 | hence "x' \<otimes> inv x = h" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 272 | by (metis assms elemrcos_carrier group.inv_solve_right mem_carrier) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 273 | thus ?thesis using h by blast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 274 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 275 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 276 | lemma (in subgroup) rcos_module_rev: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 277 | assumes "group G" "x \<in> carrier G" "x' \<in> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 278 | and "(x' \<otimes> inv x) \<in> H" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 279 | shows "x' \<in> H #> x" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 280 | proof - | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 281 | obtain h where h: "h \<in> H" "x' \<otimes> inv x = h" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 282 | using assms(4) unfolding r_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 283 | hence "x' = h \<otimes> x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 284 | by (metis assms group.inv_solve_right mem_carrier) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 285 | thus ?thesis using h unfolding r_coset_def by blast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 286 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 287 | |
| 61382 | 288 | text \<open>Module property of right cosets\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 289 | lemma (in subgroup) rcos_module: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 290 | assumes "group G" "x \<in> carrier G" "x' \<in> carrier G" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 291 | shows "(x' \<in> H #> x) = (x' \<otimes> inv x \<in> H)" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 292 | using rcos_module_rev rcos_module_imp assms by blast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 293 | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 294 | text \<open>Right cosets are subsets of the carrier.\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 295 | lemma (in subgroup) rcosets_carrier: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 296 | assumes "group G" "X \<in> rcosets H" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 297 | shows "X \<subseteq> carrier G" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 298 | using assms elemrcos_carrier singletonD | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 299 | subset_eq unfolding RCOSETS_def by force | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 300 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 301 | |
| 61382 | 302 | text \<open>Multiplication of general subsets\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 303 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 304 | lemma (in comm_group) mult_subgroups: | 
| 68604 | 305 | assumes HG: "subgroup H G" and KG: "subgroup K G" | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 306 | shows "subgroup (H <#> K) G" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 307 | proof (rule subgroup.intro) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 308 | show "H <#> K \<subseteq> carrier G" | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68445diff
changeset | 309 | by (simp add: setmult_subset_G assms subgroup.subset) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 310 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 311 | have "\<one> \<otimes> \<one> \<in> H <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 312 | unfolding set_mult_def using assms subgroup.one_closed by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 313 | thus "\<one> \<in> H <#> K" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 314 | next | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 315 | show "\<And>x. x \<in> H <#> K \<Longrightarrow> inv x \<in> H <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 316 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 317 | fix x assume "x \<in> H <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 318 | then obtain h k where hk: "h \<in> H" "k \<in> K" "x = h \<otimes> k" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 319 | unfolding set_mult_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 320 | hence "inv x = (inv k) \<otimes> (inv h)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 321 | by (meson inv_mult_group assms subgroup.mem_carrier) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 322 | hence "inv x = (inv h) \<otimes> (inv k)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 323 | by (metis hk inv_mult assms subgroup.mem_carrier) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 324 | thus "inv x \<in> H <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 325 | unfolding set_mult_def using hk assms | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 326 | by (metis (no_types, lifting) UN_iff singletonI subgroup_def) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 327 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 328 | next | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 329 | show "\<And>x y. x \<in> H <#> K \<Longrightarrow> y \<in> H <#> K \<Longrightarrow> x \<otimes> y \<in> H <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 330 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 331 | fix x y assume "x \<in> H <#> K" "y \<in> H <#> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 332 | then obtain h1 k1 h2 k2 where h1k1: "h1 \<in> H" "k1 \<in> K" "x = h1 \<otimes> k1" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 333 | and h2k2: "h2 \<in> H" "k2 \<in> K" "y = h2 \<otimes> k2" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 334 | unfolding set_mult_def by blast | 
| 68604 | 335 | with KG HG have carr: "k1 \<in> carrier G" "h1 \<in> carrier G" "k2 \<in> carrier G" "h2 \<in> carrier G" | 
| 336 | by (meson subgroup.mem_carrier)+ | |
| 337 | have "x \<otimes> y = (h1 \<otimes> k1) \<otimes> (h2 \<otimes> k2)" | |
| 338 | using h1k1 h2k2 by simp | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 339 | also have " ... = h1 \<otimes> (k1 \<otimes> h2) \<otimes> k2" | 
| 68604 | 340 | by (simp add: carr comm_groupE(3) comm_group_axioms) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 341 | also have " ... = h1 \<otimes> (h2 \<otimes> k1) \<otimes> k2" | 
| 68604 | 342 | by (simp add: carr m_comm) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 343 | finally have "x \<otimes> y = (h1 \<otimes> h2) \<otimes> (k1 \<otimes> k2)" | 
| 68604 | 344 | by (simp add: carr comm_groupE(3) comm_group_axioms) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 345 | thus "x \<otimes> y \<in> H <#> K" unfolding set_mult_def | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 346 | using subgroup.m_closed[OF assms(1) h1k1(1) h2k2(1)] | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 347 | subgroup.m_closed[OF assms(2) h1k1(2) h2k2(2)] by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 348 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 349 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 350 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 351 | lemma (in subgroup) lcos_module_rev: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 352 | assumes "group G" "x \<in> carrier G" "x' \<in> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 353 | and "(inv x \<otimes> x') \<in> H" | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 354 | shows "x' \<in> x <# H" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 355 | proof - | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 356 | obtain h where h: "h \<in> H" "inv x \<otimes> x' = h" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 357 | using assms(4) unfolding l_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 358 | hence "x' = x \<otimes> h" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 359 | by (metis assms group.inv_solve_left mem_carrier) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 360 | thus ?thesis using h unfolding l_coset_def by blast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 361 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 362 | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 363 | |
| 61382 | 364 | subsection \<open>Normal subgroups\<close> | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 365 | |
| 14963 | 366 | lemma normal_imp_subgroup: "H \<lhd> G \<Longrightarrow> subgroup H G" | 
| 367 | by (simp add: normal_def subgroup_def) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 368 | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 369 | lemma (in group) normalI: | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 370 | "subgroup H G \<Longrightarrow> (\<forall>x \<in> carrier G. H #> x = x <# H) \<Longrightarrow> H \<lhd> G" | 
| 41528 | 371 | by (simp add: normal_def normal_axioms_def is_group) | 
| 14963 | 372 | |
| 373 | lemma (in normal) inv_op_closed1: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 374 | assumes "x \<in> carrier G" and "h \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 375 | shows "(inv x) \<otimes> h \<otimes> x \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 376 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 377 | have "h \<otimes> x \<in> x <# H" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 378 | using assms coset_eq assms(1) unfolding r_coset_def by blast | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 379 | then obtain h' where "h' \<in> H" "h \<otimes> x = x \<otimes> h'" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 380 | unfolding l_coset_def by blast | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 381 | thus ?thesis by (metis assms inv_closed l_inv l_one m_assoc mem_carrier) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 382 | qed | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 383 | |
| 14963 | 384 | lemma (in normal) inv_op_closed2: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 385 | assumes "x \<in> carrier G" and "h \<in> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 386 | shows "x \<otimes> h \<otimes> (inv x) \<in> H" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 387 | using assms inv_op_closed1 by (metis inv_closed inv_inv) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 388 | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 389 | |
| 61382 | 390 | text\<open>Alternative characterization of normal subgroups\<close> | 
| 14747 | 391 | lemma (in group) normal_inv_iff: | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 392 | "(N \<lhd> G) = | 
| 67091 | 393 | (subgroup N G \<and> (\<forall>x \<in> carrier G. \<forall>h \<in> N. x \<otimes> h \<otimes> (inv x) \<in> N))" | 
| 14747 | 394 | (is "_ = ?rhs") | 
| 395 | proof | |
| 396 | assume N: "N \<lhd> G" | |
| 397 | show ?rhs | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 398 | by (blast intro: N normal.inv_op_closed2 normal_imp_subgroup) | 
| 14747 | 399 | next | 
| 400 | assume ?rhs | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 401 | hence sg: "subgroup N G" | 
| 14963 | 402 | and closed: "\<And>x. x\<in>carrier G \<Longrightarrow> \<forall>h\<in>N. x \<otimes> h \<otimes> inv x \<in> N" by auto | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 403 | hence sb: "N \<subseteq> carrier G" by (simp add: subgroup.subset) | 
| 14747 | 404 | show "N \<lhd> G" | 
| 14963 | 405 | proof (intro normalI [OF sg], simp add: l_coset_def r_coset_def, clarify) | 
| 14747 | 406 | fix x | 
| 407 | assume x: "x \<in> carrier G" | |
| 15120 | 408 |     show "(\<Union>h\<in>N. {h \<otimes> x}) = (\<Union>h\<in>N. {x \<otimes> h})"
 | 
| 14747 | 409 | proof | 
| 15120 | 410 |       show "(\<Union>h\<in>N. {h \<otimes> x}) \<subseteq> (\<Union>h\<in>N. {x \<otimes> h})"
 | 
| 14747 | 411 | proof clarify | 
| 412 | fix n | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 413 | assume n: "n \<in> N" | 
| 15120 | 414 |         show "n \<otimes> x \<in> (\<Union>h\<in>N. {x \<otimes> h})"
 | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 415 | proof | 
| 14963 | 416 | from closed [of "inv x"] | 
| 417 | show "inv x \<otimes> n \<otimes> x \<in> N" by (simp add: x n) | |
| 418 |           show "n \<otimes> x \<in> {x \<otimes> (inv x \<otimes> n \<otimes> x)}"
 | |
| 14747 | 419 | by (simp add: x n m_assoc [symmetric] sb [THEN subsetD]) | 
| 420 | qed | |
| 421 | qed | |
| 422 | next | |
| 15120 | 423 |       show "(\<Union>h\<in>N. {x \<otimes> h}) \<subseteq> (\<Union>h\<in>N. {h \<otimes> x})"
 | 
| 14747 | 424 | proof clarify | 
| 425 | fix n | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 426 | assume n: "n \<in> N" | 
| 15120 | 427 |         show "x \<otimes> n \<in> (\<Union>h\<in>N. {h \<otimes> x})"
 | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 428 | proof | 
| 14963 | 429 | show "x \<otimes> n \<otimes> inv x \<in> N" by (simp add: x n closed) | 
| 430 |           show "x \<otimes> n \<in> {x \<otimes> n \<otimes> inv x \<otimes> x}"
 | |
| 14747 | 431 | by (simp add: x n m_assoc sb [THEN subsetD]) | 
| 432 | qed | |
| 433 | qed | |
| 434 | qed | |
| 435 | qed | |
| 436 | qed | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 437 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 438 | corollary (in group) normal_invI: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 439 | assumes "subgroup N G" and "\<And>x h. \<lbrakk> x \<in> carrier G; h \<in> N \<rbrakk> \<Longrightarrow> x \<otimes> h \<otimes> inv x \<in> N" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 440 | shows "N \<lhd> G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 441 | using assms normal_inv_iff by blast | 
| 14963 | 442 | |
| 68687 | 443 | lemma (in group) one_is_normal: "{\<one>} \<lhd> G"
 | 
| 444 | proof(intro normal_invI) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 445 |   show "subgroup {\<one>} G"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 446 | by (simp add: subgroup_def) | 
| 68687 | 447 | qed simp | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 448 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 449 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 450 | subsection\<open>More Properties of Left Cosets\<close> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 451 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 452 | lemma (in group) l_repr_independence: | 
| 68687 | 453 | assumes "y \<in> x <# H" "x \<in> carrier G" and HG: "subgroup H G" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 454 | shows "x <# H = y <# H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 455 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 456 | obtain h' where h': "h' \<in> H" "y = x \<otimes> h'" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 457 | using assms(1) unfolding l_coset_def by blast | 
| 68604 | 458 | hence "x \<otimes> h = y \<otimes> ((inv h') \<otimes> h)" if "h \<in> H" for h | 
| 459 | proof - | |
| 68687 | 460 | have "h' \<in> carrier G" | 
| 461 | by (meson HG h'(1) subgroup.mem_carrier) | |
| 462 | moreover have "h \<in> carrier G" | |
| 463 | by (meson HG subgroup.mem_carrier that) | |
| 464 | ultimately show ?thesis | |
| 465 | by (metis assms(2) h'(2) inv_closed inv_solve_right m_assoc m_closed) | |
| 68604 | 466 | qed | 
| 68687 | 467 | hence "\<And>xh. xh \<in> x <# H \<Longrightarrow> xh \<in> y <# H" | 
| 468 | unfolding l_coset_def by (metis (no_types, lifting) UN_iff HG h'(1) subgroup_def) | |
| 469 | moreover have "\<And>h. h \<in> H \<Longrightarrow> y \<otimes> h = x \<otimes> (h' \<otimes> h)" | |
| 470 | using h' by (meson assms(2) HG m_assoc subgroup.mem_carrier) | |
| 471 | hence "\<And>yh. yh \<in> y <# H \<Longrightarrow> yh \<in> x <# H" | |
| 472 | unfolding l_coset_def using subgroup.m_closed[OF HG h'(1)] by blast | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 473 | ultimately show ?thesis by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 474 | qed | 
| 14803 | 475 | |
| 14747 | 476 | lemma (in group) lcos_m_assoc: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 477 | "\<lbrakk> M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G \<rbrakk> \<Longrightarrow> g <# (h <# M) = (g \<otimes> h) <# M" | 
| 14747 | 478 | by (force simp add: l_coset_def m_assoc) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 479 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 480 | lemma (in group) lcos_mult_one: "M \<subseteq> carrier G \<Longrightarrow> \<one> <# M = M" | 
| 14747 | 481 | by (force simp add: l_coset_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 482 | |
| 14747 | 483 | lemma (in group) l_coset_subset_G: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 484 | "\<lbrakk> H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> x <# H \<subseteq> carrier G" | 
| 14747 | 485 | by (auto simp add: l_coset_def subsetD) | 
| 486 | ||
| 487 | lemma (in group) l_coset_carrier: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 488 | "\<lbrakk> y \<in> x <# H; x \<in> carrier G; subgroup H G \<rbrakk> \<Longrightarrow> y \<in> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 489 | by (auto simp add: l_coset_def m_assoc subgroup.subset [THEN subsetD] subgroup.m_closed) | 
| 14530 | 490 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 491 | lemma (in group) l_coset_swap: | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 492 | assumes "y \<in> x <# H" "x \<in> carrier G" "subgroup H G" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 493 | shows "x \<in> y <# H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 494 | using assms(2) l_repr_independence[OF assms] subgroup.one_closed[OF assms(3)] | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 495 | unfolding l_coset_def by fastforce | 
| 14530 | 496 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 497 | lemma (in group) subgroup_mult_id: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 498 | assumes "subgroup H G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 499 | shows "H <#> H = H" | 
| 14666 | 500 | proof | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 501 | show "H <#> H \<subseteq> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 502 | unfolding set_mult_def using subgroup.m_closed[OF assms] by (simp add: UN_subset_iff) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 503 | show "H \<subseteq> H <#> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 504 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 505 | fix x assume x: "x \<in> H" thus "x \<in> H <#> H" unfolding set_mult_def | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 506 | using subgroup.m_closed[OF assms subgroup.one_closed[OF assms] x] subgroup.one_closed[OF assms] | 
| 68604 | 507 | using assms subgroup.mem_carrier by force | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 508 | qed | 
| 14530 | 509 | qed | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 510 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 511 | |
| 63167 | 512 | subsubsection \<open>Set of Inverses of an \<open>r_coset\<close>.\<close> | 
| 14666 | 513 | |
| 14963 | 514 | lemma (in normal) rcos_inv: | 
| 515 | assumes x: "x \<in> carrier G" | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 516 | shows "set_inv (H #> x) = H #> (inv x)" | 
| 14963 | 517 | proof (simp add: r_coset_def SET_INV_def x inv_mult_group, safe) | 
| 518 | fix h | |
| 41528 | 519 | assume h: "h \<in> H" | 
| 15120 | 520 |   show "inv x \<otimes> inv h \<in> (\<Union>j\<in>H. {j \<otimes> inv x})"
 | 
| 14963 | 521 | proof | 
| 522 | show "inv x \<otimes> inv h \<otimes> x \<in> H" | |
| 41528 | 523 | by (simp add: inv_op_closed1 h x) | 
| 14963 | 524 |     show "inv x \<otimes> inv h \<in> {inv x \<otimes> inv h \<otimes> x \<otimes> inv x}"
 | 
| 41528 | 525 | by (simp add: h x m_assoc) | 
| 14963 | 526 | qed | 
| 15120 | 527 |   show "h \<otimes> inv x \<in> (\<Union>j\<in>H. {inv x \<otimes> inv j})"
 | 
| 14963 | 528 | proof | 
| 529 | show "x \<otimes> inv h \<otimes> inv x \<in> H" | |
| 41528 | 530 | by (simp add: inv_op_closed2 h x) | 
| 14963 | 531 |     show "h \<otimes> inv x \<in> {inv x \<otimes> inv (x \<otimes> inv h \<otimes> inv x)}"
 | 
| 41528 | 532 | by (simp add: h x m_assoc [symmetric] inv_mult_group) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 533 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 534 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 535 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 536 | |
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 537 | subsubsection \<open>Theorems for \<open><#>\<close> with \<open>#>\<close> or \<open><#\<close>.\<close> | 
| 14666 | 538 | |
| 14747 | 539 | lemma (in group) setmult_rcos_assoc: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 540 | "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk> \<Longrightarrow> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 541 | H <#> (K #> x) = (H <#> K) #> x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 542 |   using set_mult_assoc[of H K "{x}"] by (simp add: r_coset_eq_set_mult)
 | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 543 | |
| 14747 | 544 | lemma (in group) rcos_assoc_lcos: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 545 | "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk> \<Longrightarrow> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 546 | (H #> x) <#> K = H <#> (x <# K)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 547 |   using set_mult_assoc[of H "{x}" K]
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 548 | by (simp add: l_coset_eq_set_mult r_coset_eq_set_mult) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 549 | |
| 14963 | 550 | lemma (in normal) rcos_mult_step1: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 551 | "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> \<Longrightarrow> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 552 | (H #> x) <#> (H #> y) = (H <#> (x <# H)) #> y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 553 | by (simp add: setmult_rcos_assoc r_coset_subset_G | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 554 | subset l_coset_subset_G rcos_assoc_lcos) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 555 | |
| 14963 | 556 | lemma (in normal) rcos_mult_step2: | 
| 557 | "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> | |
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 558 | \<Longrightarrow> (H <#> (x <# H)) #> y = (H <#> (H #> x)) #> y" | 
| 14963 | 559 | by (insert coset_eq, simp add: normal_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 560 | |
| 14963 | 561 | lemma (in normal) rcos_mult_step3: | 
| 562 | "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> | |
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 563 | \<Longrightarrow> (H <#> (H #> x)) #> y = H #> (x \<otimes> y)" | 
| 14963 | 564 | by (simp add: setmult_rcos_assoc coset_mult_assoc | 
| 41528 | 565 | subgroup_mult_id normal.axioms subset normal_axioms) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 566 | |
| 14963 | 567 | lemma (in normal) rcos_sum: | 
| 568 | "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> | |
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 569 | \<Longrightarrow> (H #> x) <#> (H #> y) = H #> (x \<otimes> y)" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 570 | by (simp add: rcos_mult_step1 rcos_mult_step2 rcos_mult_step3) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 571 | |
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 572 | lemma (in normal) rcosets_mult_eq: "M \<in> rcosets H \<Longrightarrow> H <#> M = M" | 
| 63167 | 573 | \<comment> \<open>generalizes \<open>subgroup_mult_id\<close>\<close> | 
| 14963 | 574 | by (auto simp add: RCOSETS_def subset | 
| 41528 | 575 | setmult_rcos_assoc subgroup_mult_id normal.axioms normal_axioms) | 
| 14963 | 576 | |
| 577 | ||
| 61382 | 578 | subsubsection\<open>An Equivalence Relation\<close> | 
| 14963 | 579 | |
| 35847 | 580 | definition | 
| 581 |   r_congruent :: "[('a,'b)monoid_scheme, 'a set] \<Rightarrow> ('a*'a)set"  ("rcong\<index> _")
 | |
| 67091 | 582 |   where "rcong\<^bsub>G\<^esub> H = {(x,y). x \<in> carrier G \<and> y \<in> carrier G \<and> inv\<^bsub>G\<^esub> x \<otimes>\<^bsub>G\<^esub> y \<in> H}"
 | 
| 14963 | 583 | |
| 584 | ||
| 585 | lemma (in subgroup) equiv_rcong: | |
| 27611 | 586 | assumes "group G" | 
| 14963 | 587 | shows "equiv (carrier G) (rcong H)" | 
| 27611 | 588 | proof - | 
| 29237 | 589 | interpret group G by fact | 
| 27611 | 590 | show ?thesis | 
| 40815 | 591 | proof (intro equivI) | 
| 30198 | 592 | show "refl_on (carrier G) (rcong H)" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 593 | by (auto simp add: r_congruent_def refl_on_def) | 
| 27611 | 594 | next | 
| 595 | show "sym (rcong H)" | |
| 596 | proof (simp add: r_congruent_def sym_def, clarify) | |
| 597 | fix x y | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 598 | assume [simp]: "x \<in> carrier G" "y \<in> carrier G" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31727diff
changeset | 599 | and "inv x \<otimes> y \<in> H" | 
| 46721 | 600 | hence "inv (inv x \<otimes> y) \<in> H" by simp | 
| 27611 | 601 | thus "inv y \<otimes> x \<in> H" by (simp add: inv_mult_group) | 
| 602 | qed | |
| 603 | next | |
| 604 | show "trans (rcong H)" | |
| 605 | proof (simp add: r_congruent_def trans_def, clarify) | |
| 606 | fix x y z | |
| 607 | assume [simp]: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31727diff
changeset | 608 | and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H" | 
| 27611 | 609 | hence "(inv x \<otimes> y) \<otimes> (inv y \<otimes> z) \<in> H" by simp | 
| 27698 | 610 | hence "inv x \<otimes> (y \<otimes> inv y) \<otimes> z \<in> H" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 611 | by (simp add: m_assoc del: r_inv Units_r_inv) | 
| 27611 | 612 | thus "inv x \<otimes> z \<in> H" by simp | 
| 613 | qed | |
| 14963 | 614 | qed | 
| 615 | qed | |
| 616 | ||
| 63167 | 617 | text\<open>Equivalence classes of \<open>rcong\<close> correspond to left cosets. | 
| 14963 | 618 | Was there a mistake in the definitions? I'd have expected them to | 
| 61382 | 619 | correspond to right cosets.\<close> | 
| 14963 | 620 | |
| 621 | (* CB: This is correct, but subtle. | |
| 622 | We call H #> a the right coset of a relative to H. According to | |
| 623 | Jacobson, this is what the majority of group theory literature does. | |
| 624 | He then defines the notion of congruence relation ~ over monoids as | |
| 625 | equivalence relation with a ~ a' & b ~ b' \<Longrightarrow> a*b ~ a'*b'. | |
| 626 | Our notion of right congruence induced by K: rcong K appears only in | |
| 627 | the context where K is a normal subgroup. Jacobson doesn't name it. | |
| 628 | But in this context left and right cosets are identical. | |
| 629 | *) | |
| 630 | ||
| 631 | lemma (in subgroup) l_coset_eq_rcong: | |
| 27611 | 632 | assumes "group G" | 
| 14963 | 633 | assumes a: "a \<in> carrier G" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 634 |   shows "a <# H = (rcong H) `` {a}"
 | 
| 27611 | 635 | proof - | 
| 29237 | 636 | interpret group G by fact | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 637 | show ?thesis by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a ) | 
| 27611 | 638 | qed | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 639 | |
| 35849 | 640 | |
| 61382 | 641 | subsubsection\<open>Two Distinct Right Cosets are Disjoint\<close> | 
| 14803 | 642 | |
| 643 | lemma (in group) rcos_equation: | |
| 27611 | 644 | assumes "subgroup H G" | 
| 645 | assumes p: "ha \<otimes> a = h \<otimes> b" "a \<in> carrier G" "b \<in> carrier G" "h \<in> H" "ha \<in> H" "hb \<in> H" | |
| 646 |   shows "hb \<otimes> a \<in> (\<Union>h\<in>H. {h \<otimes> b})"
 | |
| 647 | proof - | |
| 29237 | 648 | interpret subgroup H G by fact | 
| 68687 | 649 | from p show ?thesis | 
| 650 | by (rule_tac UN_I [of "hb \<otimes> ((inv ha) \<otimes> h)"]) (auto simp: inv_solve_left m_assoc) | |
| 27611 | 651 | qed | 
| 14803 | 652 | |
| 653 | lemma (in group) rcos_disjoint: | |
| 27611 | 654 | assumes "subgroup H G" | 
| 655 | assumes p: "a \<in> rcosets H" "b \<in> rcosets H" "a\<noteq>b" | |
| 656 |   shows "a \<inter> b = {}"
 | |
| 657 | proof - | |
| 29237 | 658 | interpret subgroup H G by fact | 
| 41528 | 659 | from p show ?thesis | 
| 68687 | 660 | unfolding RCOSETS_def r_coset_def | 
| 661 | by (blast intro: rcos_equation assms sym) | |
| 27611 | 662 | qed | 
| 14803 | 663 | |
| 35849 | 664 | |
| 63167 | 665 | subsection \<open>Further lemmas for \<open>r_congruent\<close>\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 666 | |
| 61382 | 667 | text \<open>The relation is a congruence\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 668 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 669 | lemma (in normal) congruent_rcong: | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 670 | shows "congruent2 (rcong H) (rcong H) (\<lambda>a b. a \<otimes> b <# H)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 671 | proof (intro congruent2I[of "carrier G" _ "carrier G" _] equiv_rcong is_group) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 672 | fix a b c | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 673 | assume abrcong: "(a, b) \<in> rcong H" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 674 | and ccarr: "c \<in> carrier G" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 675 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 676 | from abrcong | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 677 | have acarr: "a \<in> carrier G" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 678 | and bcarr: "b \<in> carrier G" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 679 | and abH: "inv a \<otimes> b \<in> H" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 680 | unfolding r_congruent_def | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 681 | by fast+ | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 682 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 683 | note carr = acarr bcarr ccarr | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 684 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 685 | from ccarr and abH | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 686 | have "inv c \<otimes> (inv a \<otimes> b) \<otimes> c \<in> H" by (rule inv_op_closed1) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 687 | moreover | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 688 | from carr and inv_closed | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 689 | have "inv c \<otimes> (inv a \<otimes> b) \<otimes> c = (inv c \<otimes> inv a) \<otimes> (b \<otimes> c)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 690 | by (force cong: m_assoc) | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 691 | moreover | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 692 | from carr and inv_closed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 693 | have "\<dots> = (inv (a \<otimes> c)) \<otimes> (b \<otimes> c)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 694 | by (simp add: inv_mult_group) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 695 | ultimately | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 696 | have "(inv (a \<otimes> c)) \<otimes> (b \<otimes> c) \<in> H" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 697 | from carr and this | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 698 | have "(b \<otimes> c) \<in> (a \<otimes> c) <# H" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 699 | by (simp add: lcos_module_rev[OF is_group]) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 700 | from carr and this and is_subgroup | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 701 | show "(a \<otimes> c) <# H = (b \<otimes> c) <# H" by (intro l_repr_independence, simp+) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 702 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 703 | fix a b c | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 704 | assume abrcong: "(a, b) \<in> rcong H" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 705 | and ccarr: "c \<in> carrier G" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 706 | |
| 46721 | 707 | from ccarr have "c \<in> Units G" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 708 | hence cinvc_one: "inv c \<otimes> c = \<one>" by (rule Units_l_inv) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 709 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 710 | from abrcong | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 711 | have acarr: "a \<in> carrier G" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 712 | and bcarr: "b \<in> carrier G" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 713 | and abH: "inv a \<otimes> b \<in> H" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 714 | by (unfold r_congruent_def, fast+) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 715 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 716 | note carr = acarr bcarr ccarr | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 717 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 718 | from carr and inv_closed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 719 | have "inv a \<otimes> b = inv a \<otimes> (\<one> \<otimes> b)" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 720 | also from carr and inv_closed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 721 | have "\<dots> = inv a \<otimes> (inv c \<otimes> c) \<otimes> b" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 722 | also from carr and inv_closed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 723 | have "\<dots> = (inv a \<otimes> inv c) \<otimes> (c \<otimes> b)" by (force cong: m_assoc) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 724 | also from carr and inv_closed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 725 | have "\<dots> = inv (c \<otimes> a) \<otimes> (c \<otimes> b)" by (simp add: inv_mult_group) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 726 | finally | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 727 | have "inv a \<otimes> b = inv (c \<otimes> a) \<otimes> (c \<otimes> b)" . | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 728 | from abH and this | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 729 | have "inv (c \<otimes> a) \<otimes> (c \<otimes> b) \<in> H" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 730 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 731 | from carr and this | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 732 | have "(c \<otimes> b) \<in> (c \<otimes> a) <# H" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 733 | by (simp add: lcos_module_rev[OF is_group]) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 734 | from carr and this and is_subgroup | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 735 | show "(c \<otimes> a) <# H = (c \<otimes> b) <# H" by (intro l_repr_independence, simp+) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 736 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
19931diff
changeset | 737 | |
| 14803 | 738 | |
| 61382 | 739 | subsection \<open>Order of a Group and Lagrange's Theorem\<close> | 
| 14803 | 740 | |
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 741 | definition | 
| 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 742 |   order :: "('a, 'b) monoid_scheme \<Rightarrow> nat"
 | 
| 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 743 | where "order S = card (carrier S)" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 744 | |
| 61628 | 745 | lemma (in monoid) order_gt_0_iff_finite: "0 < order G \<longleftrightarrow> finite (carrier G)" | 
| 746 | by(auto simp add: order_def card_gt_0_iff) | |
| 747 | ||
| 14963 | 748 | lemma (in group) rcosets_part_G: | 
| 27611 | 749 | assumes "subgroup H G" | 
| 14963 | 750 | shows "\<Union>(rcosets H) = carrier G" | 
| 27611 | 751 | proof - | 
| 29237 | 752 | interpret subgroup H G by fact | 
| 27611 | 753 | show ?thesis | 
| 68687 | 754 | unfolding RCOSETS_def r_coset_def by auto | 
| 27611 | 755 | qed | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 756 | |
| 14747 | 757 | lemma (in group) cosets_finite: | 
| 14963 | 758 | "\<lbrakk>c \<in> rcosets H; H \<subseteq> carrier G; finite (carrier G)\<rbrakk> \<Longrightarrow> finite c" | 
| 68687 | 759 | unfolding RCOSETS_def | 
| 760 | by (auto simp add: r_coset_subset_G [THEN finite_subset]) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 761 | |
| 63167 | 762 | text\<open>The next two lemmas support the proof of \<open>card_cosets_equal\<close>.\<close> | 
| 14747 | 763 | lemma (in group) inj_on_f: | 
| 68687 | 764 | assumes "H \<subseteq> carrier G" and a: "a \<in> carrier G" | 
| 765 | shows "inj_on (\<lambda>y. y \<otimes> inv a) (H #> a)" | |
| 766 | proof | |
| 767 | fix x y | |
| 768 | assume "x \<in> H #> a" "y \<in> H #> a" and xy: "x \<otimes> inv a = y \<otimes> inv a" | |
| 769 | then have "x \<in> carrier G" "y \<in> carrier G" | |
| 770 | using assms r_coset_subset_G by blast+ | |
| 771 | with xy a show "x = y" | |
| 772 | by auto | |
| 773 | qed | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 774 | |
| 14747 | 775 | lemma (in group) inj_on_g: | 
| 14963 | 776 | "\<lbrakk>H \<subseteq> carrier G; a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. y \<otimes> a) H" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 777 | by (force simp add: inj_on_def subsetD) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 778 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 779 | (* ************************************************************************** *) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 780 | |
| 14747 | 781 | lemma (in group) card_cosets_equal: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 782 | assumes "R \<in> rcosets H" "H \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 783 | shows "\<exists>f. bij_betw f H R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 784 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 785 | obtain g where g: "g \<in> carrier G" "R = H #> g" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 786 | using assms(1) unfolding RCOSETS_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 787 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 788 | let ?f = "\<lambda>h. h \<otimes> g" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 789 | have "\<And>r. r \<in> R \<Longrightarrow> \<exists>h \<in> H. ?f h = r" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 790 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 791 | fix r assume "r \<in> R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 792 | then obtain h where "h \<in> H" "r = h \<otimes> g" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 793 | using g unfolding r_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 794 | thus "\<exists>h \<in> H. ?f h = r" by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 795 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 796 | hence "R \<subseteq> ?f ` H" by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 797 | moreover have "?f ` H \<subseteq> R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 798 | using g unfolding r_coset_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 799 | ultimately show ?thesis using inj_on_g unfolding bij_betw_def | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 800 | using assms(2) g(1) by auto | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 801 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 802 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 803 | corollary (in group) card_rcosets_equal: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 804 | assumes "R \<in> rcosets H" "H \<subseteq> carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 805 | shows "card H = card R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 806 | using card_cosets_equal assms bij_betw_same_card by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 807 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 808 | corollary (in group) rcosets_finite: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 809 | assumes "R \<in> rcosets H" "H \<subseteq> carrier G" "finite H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 810 | shows "finite R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 811 | using card_cosets_equal assms bij_betw_finite is_group by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 812 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 813 | (* ************************************************************************** *) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 814 | |
| 14963 | 815 | lemma (in group) rcosets_subset_PowG: | 
| 816 | "subgroup H G \<Longrightarrow> rcosets H \<subseteq> Pow(carrier G)" | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 817 | using rcosets_part_G by auto | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 818 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 819 | proposition (in group) lagrange_finite: | 
| 68687 | 820 | assumes "finite(carrier G)" and HG: "subgroup H G" | 
| 821 | shows "card(rcosets H) * card(H) = order(G)" | |
| 822 | proof - | |
| 823 | have "card H * card (rcosets H) = card (\<Union>(rcosets H))" | |
| 824 | proof (rule card_partition) | |
| 825 |     show "\<And>c1 c2. \<lbrakk>c1 \<in> rcosets H; c2 \<in> rcosets H; c1 \<noteq> c2\<rbrakk> \<Longrightarrow> c1 \<inter> c2 = {}"
 | |
| 826 | using HG rcos_disjoint by auto | |
| 827 | qed (auto simp: assms finite_UnionD rcosets_part_G card_rcosets_equal subgroup.subset) | |
| 828 | then show ?thesis | |
| 829 | by (simp add: HG mult.commute order_def rcosets_part_G) | |
| 830 | qed | |
| 14803 | 831 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 832 | theorem (in group) lagrange: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 833 | assumes "subgroup H G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 834 | shows "card (rcosets H) * card H = order G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 835 | proof (cases "finite (carrier G)") | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 836 | case True thus ?thesis using lagrange_finite assms by simp | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 837 | next | 
| 68687 | 838 | case False | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 839 | thus ?thesis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 840 | proof (cases "finite H") | 
| 68687 | 841 | case False thus ?thesis using \<open>infinite (carrier G)\<close> by (simp add: order_def) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 842 | next | 
| 68687 | 843 | case True | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 844 | have "infinite (rcosets H)" | 
| 68687 | 845 | proof | 
| 846 | assume "finite (rcosets H)" | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 847 | hence finite_rcos: "finite (rcosets H)" by simp | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 848 | hence "card (\<Union>(rcosets H)) = (\<Sum>R\<in>(rcosets H). card R)" | 
| 68687 | 849 | using card_Union_disjoint[of "rcosets H"] \<open>finite H\<close> rcos_disjoint[OF assms(1)] | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68445diff
changeset | 850 | rcosets_finite[where ?H = H] by (simp add: assms subgroup.subset) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 851 | hence "order G = (\<Sum>R\<in>(rcosets H). card R)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 852 | by (simp add: assms order_def rcosets_part_G) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 853 | hence "order G = (\<Sum>R\<in>(rcosets H). card H)" | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68445diff
changeset | 854 | using card_rcosets_equal by (simp add: assms subgroup.subset) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 855 | hence "order G = (card H) * (card (rcosets H))" by simp | 
| 68687 | 856 | hence "order G \<noteq> 0" using finite_rcos \<open>finite H\<close> assms ex_in_conv | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 857 | rcosets_part_G subgroup.one_closed by fastforce | 
| 68687 | 858 | thus False using \<open>infinite (carrier G)\<close> order_gt_0_iff_finite by blast | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 859 | qed | 
| 68687 | 860 | thus ?thesis using \<open>infinite (carrier G)\<close> by (simp add: order_def) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 861 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 862 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 863 | |
| 14803 | 864 | |
| 61382 | 865 | subsection \<open>Quotient Groups: Factorization of a Group\<close> | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 866 | |
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 867 | definition | 
| 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 868 |   FactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "Mod" 65)
 | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67091diff
changeset | 869 | \<comment> \<open>Actually defined for groups rather than monoids\<close> | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 870 | where "FactGroup G H = \<lparr>carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>" | 
| 14747 | 871 | |
| 14963 | 872 | lemma (in normal) setmult_closed: | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 873 | "\<lbrakk>K1 \<in> rcosets H; K2 \<in> rcosets H\<rbrakk> \<Longrightarrow> K1 <#> K2 \<in> rcosets H" | 
| 14963 | 874 | by (auto simp add: rcos_sum RCOSETS_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 875 | |
| 14963 | 876 | lemma (in normal) setinv_closed: | 
| 877 | "K \<in> rcosets H \<Longrightarrow> set_inv K \<in> rcosets H" | |
| 878 | by (auto simp add: rcos_inv RCOSETS_def) | |
| 13889 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 879 | |
| 14963 | 880 | lemma (in normal) rcosets_assoc: | 
| 881 | "\<lbrakk>M1 \<in> rcosets H; M2 \<in> rcosets H; M3 \<in> rcosets H\<rbrakk> | |
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 882 | \<Longrightarrow> M1 <#> M2 <#> M3 = M1 <#> (M2 <#> M3)" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 883 | by (simp add: group.set_mult_assoc is_group rcosets_carrier) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 884 | |
| 14963 | 885 | lemma (in subgroup) subgroup_in_rcosets: | 
| 27611 | 886 | assumes "group G" | 
| 14963 | 887 | shows "H \<in> rcosets H" | 
| 13889 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 888 | proof - | 
| 29237 | 889 | interpret group G by fact | 
| 26203 | 890 | from _ subgroup_axioms have "H #> \<one> = H" | 
| 23350 | 891 | by (rule coset_join2) auto | 
| 13889 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 892 | then show ?thesis | 
| 14963 | 893 | by (auto simp add: RCOSETS_def) | 
| 13889 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 894 | qed | 
| 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 895 | |
| 14963 | 896 | lemma (in normal) rcosets_inv_mult_group_eq: | 
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 897 | "M \<in> rcosets H \<Longrightarrow> set_inv M <#> M = H" | 
| 41528 | 898 | by (auto simp add: RCOSETS_def rcos_inv rcos_sum subgroup.subset normal.axioms normal_axioms) | 
| 13889 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 899 | |
| 14963 | 900 | theorem (in normal) factorgroup_is_group: | 
| 901 | "group (G Mod H)" | |
| 68687 | 902 | unfolding FactGroup_def | 
| 903 | apply (rule groupI) | |
| 14747 | 904 | apply (simp add: setmult_closed) | 
| 14963 | 905 | apply (simp add: normal_imp_subgroup subgroup_in_rcosets [OF is_group]) | 
| 906 | apply (simp add: restrictI setmult_closed rcosets_assoc) | |
| 13889 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 907 | apply (simp add: normal_imp_subgroup | 
| 14963 | 908 | subgroup_in_rcosets rcosets_mult_eq) | 
| 909 | apply (auto dest: rcosets_inv_mult_group_eq simp add: setinv_closed) | |
| 13889 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 910 | done | 
| 
6676ac2527fa
Fixed Coset.thy (proved theorem factorgroup_is_group).
 ballarin parents: 
13870diff
changeset | 911 | |
| 65035 
b46fe5138cb0
backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
 haftmann parents: 
64587diff
changeset | 912 | lemma mult_FactGroup [simp]: "X \<otimes>\<^bsub>(G Mod H)\<^esub> X' = X <#>\<^bsub>G\<^esub> X'" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 913 | by (simp add: FactGroup_def) | 
| 14803 | 914 | |
| 14963 | 915 | lemma (in normal) inv_FactGroup: | 
| 68687 | 916 | assumes "X \<in> carrier (G Mod H)" | 
| 917 | shows "inv\<^bsub>G Mod H\<^esub> X = set_inv X" | |
| 918 | proof - | |
| 919 | have X: "X \<in> rcosets H" | |
| 920 | using assms by (simp add: FactGroup_def) | |
| 921 | moreover have "set_inv X <#> X = H" | |
| 922 | using X by (simp add: normal.rcosets_inv_mult_group_eq normal_axioms) | |
| 923 | moreover have "Group.group (G Mod H)" | |
| 924 | using normal.factorgroup_is_group normal_axioms by blast | |
| 925 | moreover have "set_inv X \<in> rcosets H" | |
| 926 | by (simp add: \<open>X \<in> rcosets H\<close> setinv_closed) | |
| 927 | ultimately show ?thesis | |
| 928 | by (simp add: FactGroup_def group.inv_equality) | |
| 929 | qed | |
| 14747 | 930 | |
| 61382 | 931 | text\<open>The coset map is a homomorphism from @{term G} to the quotient group
 | 
| 932 |   @{term "G Mod H"}\<close>
 | |
| 14963 | 933 | lemma (in normal) r_coset_hom_Mod: | 
| 934 | "(\<lambda>a. H #> a) \<in> hom G (G Mod H)" | |
| 935 | by (auto simp add: FactGroup_def RCOSETS_def Pi_def hom_def rcos_sum) | |
| 14747 | 936 | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 937 | |
| 61382 | 938 | subsection\<open>The First Isomorphism Theorem\<close> | 
| 14803 | 939 | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 940 | text\<open>The quotient by the kernel of a homomorphism is isomorphic to the | 
| 61382 | 941 | range of that homomorphism.\<close> | 
| 14803 | 942 | |
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 943 | definition | 
| 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 944 |   kernel :: "('a, 'm) monoid_scheme \<Rightarrow> ('b, 'n) monoid_scheme \<Rightarrow>  ('a \<Rightarrow> 'b) \<Rightarrow> 'a set"
 | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67091diff
changeset | 945 | \<comment> \<open>the kernel of a homomorphism\<close> | 
| 67091 | 946 |   where "kernel G H h = {x. x \<in> carrier G \<and> h x = \<one>\<^bsub>H\<^esub>}"
 | 
| 14803 | 947 | |
| 948 | lemma (in group_hom) subgroup_kernel: "subgroup (kernel G H h) G" | |
| 68687 | 949 | by (auto simp add: kernel_def group.intro is_group intro: subgroup.intro) | 
| 14803 | 950 | |
| 61382 | 951 | text\<open>The kernel of a homomorphism is a normal subgroup\<close> | 
| 14963 | 952 | lemma (in group_hom) normal_kernel: "(kernel G H h) \<lhd> G" | 
| 68687 | 953 | apply (simp only: G.normal_inv_iff subgroup_kernel) | 
| 954 | apply (simp add: kernel_def) | |
| 955 | done | |
| 14803 | 956 | |
| 957 | lemma (in group_hom) FactGroup_nonempty: | |
| 958 | assumes X: "X \<in> carrier (G Mod kernel G H h)" | |
| 959 |   shows "X \<noteq> {}"
 | |
| 960 | proof - | |
| 961 | from X | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 962 | obtain g where "g \<in> carrier G" | 
| 14803 | 963 | and "X = kernel G H h #> g" | 
| 14963 | 964 | by (auto simp add: FactGroup_def RCOSETS_def) | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 965 | thus ?thesis | 
| 14963 | 966 | by (auto simp add: kernel_def r_coset_def image_def intro: hom_one) | 
| 14803 | 967 | qed | 
| 968 | ||
| 969 | ||
| 39910 | 970 | lemma (in group_hom) FactGroup_the_elem_mem: | 
| 14803 | 971 | assumes X: "X \<in> carrier (G Mod (kernel G H h))" | 
| 39910 | 972 | shows "the_elem (h`X) \<in> carrier H" | 
| 14803 | 973 | proof - | 
| 974 | from X | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 975 | obtain g where g: "g \<in> carrier G" | 
| 14803 | 976 | and "X = kernel G H h #> g" | 
| 14963 | 977 | by (auto simp add: FactGroup_def RCOSETS_def) | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61628diff
changeset | 978 |   hence "h ` X = {h g}" by (auto simp add: kernel_def r_coset_def g intro!: imageI)
 | 
| 14803 | 979 | thus ?thesis by (auto simp add: g) | 
| 980 | qed | |
| 981 | ||
| 982 | lemma (in group_hom) FactGroup_hom: | |
| 39910 | 983 | "(\<lambda>X. the_elem (h`X)) \<in> hom (G Mod (kernel G H h)) H" | 
| 68687 | 984 | proof - | 
| 985 | have "the_elem (h ` (X <#> X')) = the_elem (h ` X) \<otimes>\<^bsub>H\<^esub> the_elem (h ` X')" | |
| 986 | if X: "X \<in> carrier (G Mod kernel G H h)" and X': "X' \<in> carrier (G Mod kernel G H h)" for X X' | |
| 987 | proof - | |
| 988 | obtain g and g' | |
| 989 | where "g \<in> carrier G" and "g' \<in> carrier G" | |
| 990 | and "X = kernel G H h #> g" and "X' = kernel G H h #> g'" | |
| 991 | using X X' by (auto simp add: FactGroup_def RCOSETS_def) | |
| 992 | hence all: "\<forall>x\<in>X. h x = h g" "\<forall>x\<in>X'. h x = h g'" | |
| 993 | and Xsub: "X \<subseteq> carrier G" and X'sub: "X' \<subseteq> carrier G" | |
| 994 | by (force simp add: kernel_def r_coset_def image_def)+ | |
| 995 |     hence "h ` (X <#> X') = {h g \<otimes>\<^bsub>H\<^esub> h g'}" using X X'
 | |
| 996 | by (auto dest!: FactGroup_nonempty intro!: image_eqI | |
| 997 | simp add: set_mult_def | |
| 998 | subsetD [OF Xsub] subsetD [OF X'sub]) | |
| 999 | then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \<otimes>\<^bsub>H\<^esub> the_elem (h ` X')" | |
| 1000 | by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique) | |
| 1001 | qed | |
| 1002 | then show ?thesis | |
| 1003 | by (simp add: hom_def FactGroup_the_elem_mem normal.factorgroup_is_group [OF normal_kernel] group.axioms monoid.m_closed) | |
| 14803 | 1004 | qed | 
| 1005 | ||
| 14963 | 1006 | |
| 61382 | 1007 | text\<open>Lemma for the following injectivity result\<close> | 
| 14803 | 1008 | lemma (in group_hom) FactGroup_subset: | 
| 68687 | 1009 | assumes "g \<in> carrier G" "g' \<in> carrier G" "h g = h g'" | 
| 1010 | shows "kernel G H h #> g \<subseteq> kernel G H h #> g'" | |
| 1011 | unfolding kernel_def r_coset_def | |
| 1012 | proof clarsimp | |
| 1013 | fix y | |
| 1014 | assume "y \<in> carrier G" "h y = \<one>\<^bsub>H\<^esub>" | |
| 1015 | with assms show "\<exists>x. x \<in> carrier G \<and> h x = \<one>\<^bsub>H\<^esub> \<and> y \<otimes> g = x \<otimes> g'" | |
| 1016 | by (rule_tac x="y \<otimes> g \<otimes> inv g'" in exI) (auto simp: G.m_assoc) | |
| 1017 | qed | |
| 14803 | 1018 | |
| 1019 | lemma (in group_hom) FactGroup_inj_on: | |
| 39910 | 1020 | "inj_on (\<lambda>X. the_elem (h ` X)) (carrier (G Mod kernel G H h))" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1021 | proof (simp add: inj_on_def, clarify) | 
| 14803 | 1022 | fix X and X' | 
| 1023 | assume X: "X \<in> carrier (G Mod kernel G H h)" | |
| 1024 | and X': "X' \<in> carrier (G Mod kernel G H h)" | |
| 1025 | then | |
| 1026 | obtain g and g' | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1027 | where gX: "g \<in> carrier G" "g' \<in> carrier G" | 
| 14803 | 1028 | "X = kernel G H h #> g" "X' = kernel G H h #> g'" | 
| 14963 | 1029 | by (auto simp add: FactGroup_def RCOSETS_def) | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1030 | hence all: "\<forall>x\<in>X. h x = h g" "\<forall>x\<in>X'. h x = h g'" | 
| 14803 | 1031 | by (force simp add: kernel_def r_coset_def image_def)+ | 
| 39910 | 1032 | assume "the_elem (h ` X) = the_elem (h ` X')" | 
| 14803 | 1033 | hence h: "h g = h g'" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1034 | by (simp add: all FactGroup_nonempty X X' the_elem_image_unique) | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1035 | show "X=X'" by (rule equalityI) (simp_all add: FactGroup_subset h gX) | 
| 14803 | 1036 | qed | 
| 1037 | ||
| 61382 | 1038 | text\<open>If the homomorphism @{term h} is onto @{term H}, then so is the
 | 
| 1039 | homomorphism from the quotient group\<close> | |
| 14803 | 1040 | lemma (in group_hom) FactGroup_onto: | 
| 1041 | assumes h: "h ` carrier G = carrier H" | |
| 39910 | 1042 | shows "(\<lambda>X. the_elem (h ` X)) ` carrier (G Mod kernel G H h) = carrier H" | 
| 14803 | 1043 | proof | 
| 39910 | 1044 | show "(\<lambda>X. the_elem (h ` X)) ` carrier (G Mod kernel G H h) \<subseteq> carrier H" | 
| 1045 | by (auto simp add: FactGroup_the_elem_mem) | |
| 1046 | show "carrier H \<subseteq> (\<lambda>X. the_elem (h ` X)) ` carrier (G Mod kernel G H h)" | |
| 14803 | 1047 | proof | 
| 1048 | fix y | |
| 1049 | assume y: "y \<in> carrier H" | |
| 1050 | with h obtain g where g: "g \<in> carrier G" "h g = y" | |
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1051 | by (blast elim: equalityE) | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1052 |     hence "(\<Union>x\<in>kernel G H h #> g. {h x}) = {y}"
 | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1053 | by (auto simp add: y kernel_def r_coset_def) | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1054 | with g show "y \<in> (\<lambda>X. the_elem (h ` X)) ` carrier (G Mod kernel G H h)" | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61628diff
changeset | 1055 | apply (auto intro!: bexI image_eqI simp add: FactGroup_def RCOSETS_def) | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61628diff
changeset | 1056 | apply (subst the_elem_image_unique) | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61628diff
changeset | 1057 | apply auto | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61628diff
changeset | 1058 | done | 
| 14803 | 1059 | qed | 
| 1060 | qed | |
| 1061 | ||
| 1062 | ||
| 61382 | 1063 | text\<open>If @{term h} is a homomorphism from @{term G} onto @{term H}, then the
 | 
| 1064 |  quotient group @{term "G Mod (kernel G H h)"} is isomorphic to @{term H}.\<close>
 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1065 | theorem (in group_hom) FactGroup_iso_set: | 
| 14803 | 1066 | "h ` carrier G = carrier H | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1067 | \<Longrightarrow> (\<lambda>X. the_elem (h`X)) \<in> iso (G Mod (kernel G H h)) H" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1068 | by (simp add: iso_def FactGroup_hom FactGroup_inj_on bij_betw_def | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1069 | FactGroup_onto) | 
| 14803 | 1070 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1071 | corollary (in group_hom) FactGroup_iso : | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1072 | "h ` carrier G = carrier H | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1073 | \<Longrightarrow> (G Mod (kernel G H h))\<cong> H" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1074 | using FactGroup_iso_set unfolding is_iso_def by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1075 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1076 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1077 | (* Next two lemmas contributed by Paulo Emílio de Vilhena. *) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1078 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1079 | lemma (in group_hom) trivial_hom_iff: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1080 |   "(h ` (carrier G) = { \<one>\<^bsub>H\<^esub> }) = (kernel G H h = carrier G)"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1081 | unfolding kernel_def using one_closed by force | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1082 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1083 | lemma (in group_hom) trivial_ker_imp_inj: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1084 |   assumes "kernel G H h = { \<one> }"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1085 | shows "inj_on h (carrier G)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1086 | proof (rule inj_onI) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1087 | fix g1 g2 assume A: "g1 \<in> carrier G" "g2 \<in> carrier G" "h g1 = h g2" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1088 | hence "h (g1 \<otimes> (inv g2)) = \<one>\<^bsub>H\<^esub>" by simp | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1089 | hence "g1 \<otimes> (inv g2) = \<one>" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1090 | using A assms unfolding kernel_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1091 | thus "g1 = g2" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1092 | using A G.inv_equality G.inv_inv by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1093 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1094 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1095 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1096 | (* Next subsection contributed by Martin Baillon. *) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1097 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1098 | subsection \<open>Theorems about Factor Groups and Direct product\<close> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1099 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1100 | lemma (in group) DirProd_normal : | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1101 | assumes "group K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1102 | and "H \<lhd> G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1103 | and "N \<lhd> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1104 | shows "H \<times> N \<lhd> G \<times>\<times> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1105 | proof (intro group.normal_invI[OF DirProd_group[OF group_axioms assms(1)]]) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1106 | show sub : "subgroup (H \<times> N) (G \<times>\<times> K)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1107 | using DirProd_subgroups[OF group_axioms normal_imp_subgroup[OF assms(2)]assms(1) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1108 | normal_imp_subgroup[OF assms(3)]]. | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1109 | show "\<And>x h. x \<in> carrier (G\<times>\<times>K) \<Longrightarrow> h \<in> H\<times>N \<Longrightarrow> x \<otimes>\<^bsub>G\<times>\<times>K\<^esub> h \<otimes>\<^bsub>G\<times>\<times>K\<^esub> inv\<^bsub>G\<times>\<times>K\<^esub> x \<in> H\<times>N" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1110 | proof- | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1111 | fix x h assume xGK : "x \<in> carrier (G \<times>\<times> K)" and hHN : " h \<in> H \<times> N" | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68445diff
changeset | 1112 | hence hGK : "h \<in> carrier (G \<times>\<times> K)" using subgroup.subset[OF sub] by auto | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1113 | from xGK obtain x1 x2 where x1x2 :"x1 \<in> carrier G" "x2 \<in> carrier K" "x = (x1,x2)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1114 | unfolding DirProd_def by fastforce | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1115 | from hHN obtain h1 h2 where h1h2 : "h1 \<in> H" "h2 \<in> N" "h = (h1,h2)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1116 | unfolding DirProd_def by fastforce | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1117 | hence h1h2GK : "h1 \<in> carrier G" "h2 \<in> carrier K" | 
| 68687 | 1118 | using normal_imp_subgroup subgroup.subset assms by blast+ | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1119 | have "inv\<^bsub>G \<times>\<times> K\<^esub> x = (inv\<^bsub>G\<^esub> x1,inv\<^bsub>K\<^esub> x2)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1120 | using inv_DirProd[OF group_axioms assms(1) x1x2(1)x1x2(2)] x1x2 by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1121 | hence "x \<otimes>\<^bsub>G \<times>\<times> K\<^esub> h \<otimes>\<^bsub>G \<times>\<times> K\<^esub> inv\<^bsub>G \<times>\<times> K\<^esub> x = (x1 \<otimes> h1 \<otimes> inv x1,x2 \<otimes>\<^bsub>K\<^esub> h2 \<otimes>\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1122 | using h1h2 x1x2 h1h2GK by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1123 | moreover have "x1 \<otimes> h1 \<otimes> inv x1 \<in> H" "x2 \<otimes>\<^bsub>K\<^esub> h2 \<otimes>\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2 \<in> N" | 
| 68687 | 1124 | using assms x1x2 h1h2 assms by (simp_all add: normal.inv_op_closed2) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1125 | hence "(x1 \<otimes> h1 \<otimes> inv x1, x2 \<otimes>\<^bsub>K\<^esub> h2 \<otimes>\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2)\<in> H \<times> N" by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1126 | ultimately show " x \<otimes>\<^bsub>G \<times>\<times> K\<^esub> h \<otimes>\<^bsub>G \<times>\<times> K\<^esub> inv\<^bsub>G \<times>\<times> K\<^esub> x \<in> H \<times> N" by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1127 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1128 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1129 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1130 | lemma (in group) FactGroup_DirProd_multiplication_iso_set : | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1131 | assumes "group K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1132 | and "H \<lhd> G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1133 | and "N \<lhd> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1134 | shows "(\<lambda> (X, Y). X \<times> Y) \<in> iso ((G Mod H) \<times>\<times> (K Mod N)) (G \<times>\<times> K Mod H \<times> N)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1135 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1136 | proof- | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1137 | have R :"(\<lambda>(X, Y). X \<times> Y) \<in> carrier (G Mod H) \<times> carrier (K Mod N) \<rightarrow> carrier (G \<times>\<times> K Mod H \<times> N)" | 
| 68687 | 1138 | unfolding r_coset_def Sigma_def DirProd_def FactGroup_def RCOSETS_def by force | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1139 | moreover have "(\<forall>x\<in>carrier (G Mod H). \<forall>y\<in>carrier (K Mod N). \<forall>xa\<in>carrier (G Mod H). | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1140 | \<forall>ya\<in>carrier (K Mod N). (x <#> xa) \<times> (y <#>\<^bsub>K\<^esub> ya) = x \<times> y <#>\<^bsub>G \<times>\<times> K\<^esub> xa \<times> ya)" | 
| 68517 | 1141 | unfolding set_mult_def by force | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1142 | moreover have "(\<forall>x\<in>carrier (G Mod H). \<forall>y\<in>carrier (K Mod N). \<forall>xa\<in>carrier (G Mod H). | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1143 | \<forall>ya\<in>carrier (K Mod N). x \<times> y = xa \<times> ya \<longrightarrow> x = xa \<and> y = ya)" | 
| 68517 | 1144 | unfolding FactGroup_def using times_eq_iff subgroup.rcosets_non_empty | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1145 | by (metis assms(2) assms(3) normal_def partial_object.select_convs(1)) | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1146 | moreover have "(\<lambda>(X, Y). X \<times> Y) ` (carrier (G Mod H) \<times> carrier (K Mod N)) = | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1147 | carrier (G \<times>\<times> K Mod H \<times> N)" | 
| 68687 | 1148 | proof - | 
| 1149 | have 1: "\<And>x a b. \<lbrakk>a \<in> carrier (G Mod H); b \<in> carrier (K Mod N)\<rbrakk> \<Longrightarrow> a \<times> b \<in> carrier (G \<times>\<times> K Mod H \<times> N)" | |
| 1150 | using R by force | |
| 1151 | have 2: "\<And>z. z \<in> carrier (G \<times>\<times> K Mod H \<times> N) \<Longrightarrow> \<exists>x\<in>carrier (G Mod H). \<exists>y\<in>carrier (K Mod N). z = x \<times> y" | |
| 1152 | unfolding DirProd_def FactGroup_def RCOSETS_def r_coset_def by force | |
| 1153 | show ?thesis | |
| 1154 | unfolding image_def by (auto simp: intro: 1 2) | |
| 1155 | qed | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1156 | ultimately show ?thesis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1157 | unfolding iso_def hom_def bij_betw_def inj_on_def by simp | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1158 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1159 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1160 | corollary (in group) FactGroup_DirProd_multiplication_iso_1 : | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1161 | assumes "group K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1162 | and "H \<lhd> G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1163 | and "N \<lhd> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1164 | shows " ((G Mod H) \<times>\<times> (K Mod N)) \<cong> (G \<times>\<times> K Mod H \<times> N)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1165 | unfolding is_iso_def using FactGroup_DirProd_multiplication_iso_set assms by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1166 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1167 | corollary (in group) FactGroup_DirProd_multiplication_iso_2 : | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1168 | assumes "group K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1169 | and "H \<lhd> G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1170 | and "N \<lhd> K" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1171 | shows "(G \<times>\<times> K Mod H \<times> N) \<cong> ((G Mod H) \<times>\<times> (K Mod N))" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1172 | using FactGroup_DirProd_multiplication_iso_1 group.iso_sym assms | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1173 | DirProd_group[OF normal.factorgroup_is_group normal.factorgroup_is_group] | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1174 | by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1175 | |
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1176 | subsubsection "More Lemmas about set multiplication" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1177 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1178 | (*A group multiplied by a subgroup stays the same*) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1179 | lemma (in group) set_mult_carrier_idem: | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1180 | assumes "subgroup H G" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1181 | shows "(carrier G) <#> H = carrier G" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1182 | proof | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1183 | show "(carrier G)<#>H \<subseteq> carrier G" | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68445diff
changeset | 1184 | unfolding set_mult_def using subgroup.subset assms by blast | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1185 | next | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1186 | have " (carrier G) #> \<one> = carrier G" unfolding set_mult_def r_coset_def group_axioms by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1187 | moreover have "(carrier G) #> \<one> \<subseteq> (carrier G) <#> H" unfolding set_mult_def r_coset_def | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1188 | using assms subgroup.one_closed[OF assms] by blast | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1189 | ultimately show "carrier G \<subseteq> (carrier G) <#> H" by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1190 | qed | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1191 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1192 | (*Same lemma as above, but everything is included in a subgroup*) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1193 | lemma (in group) set_mult_subgroup_idem: | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1194 | assumes HG: "subgroup H G" and NG: "subgroup N (G \<lparr> carrier := H \<rparr>)" | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1195 | shows "H <#> N = H" | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1196 | using group.set_mult_carrier_idem[OF subgroup.subgroup_is_group[OF HG group_axioms] NG] by simp | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1197 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1198 | (*A normal subgroup is commutative with set_mult*) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1199 | lemma (in group) commut_normal: | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1200 | assumes "subgroup H G" and "N\<lhd>G" | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1201 | shows "H<#>N = N<#>H" | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1202 | proof- | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1203 |   have aux1: "{H <#> N} = {\<Union>h\<in>H. h <# N }" unfolding set_mult_def l_coset_def by auto
 | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1204 |   also have "... = {\<Union>h\<in>H. N #> h }" using assms normal.coset_eq subgroup.mem_carrier by fastforce
 | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1205 |   moreover have aux2: "{N <#> H} = {\<Union>h\<in>H. N #> h }"unfolding set_mult_def r_coset_def by auto
 | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1206 | ultimately show "H<#>N = N<#>H" by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1207 | qed | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1208 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1209 | (*Same lemma as above, but everything is included in a subgroup*) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1210 | lemma (in group) commut_normal_subgroup: | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1211 | assumes "subgroup H G" and "N \<lhd> (G\<lparr> carrier := H \<rparr>)" | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1212 | and "subgroup K (G \<lparr> carrier := H \<rparr>)" | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1213 | shows "K <#> N = N <#> K" | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1214 | using group.commut_normal[OF subgroup.subgroup_is_group[OF assms(1) group_axioms] assms(3,2)] by simp | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1215 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1216 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1217 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1218 | subsubsection "Lemmas about intersection and normal subgroups" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1219 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1220 | lemma (in group) normal_inter: | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1221 | assumes "subgroup H G" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1222 | and "subgroup K G" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1223 | and "H1\<lhd>G\<lparr>carrier := H\<rparr>" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1224 | shows " (H1\<inter>K)\<lhd>(G\<lparr>carrier:= (H\<inter>K)\<rparr>)" | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1225 | proof- | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1226 | define HK and H1K and GH and GHK | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1227 | where "HK = H\<inter>K" and "H1K=H1\<inter>K" and "GH =G\<lparr>carrier := H\<rparr>" and "GHK = (G\<lparr>carrier:= (H\<inter>K)\<rparr>)" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1228 | show "H1K\<lhd>GHK" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1229 | proof (intro group.normal_invI[of GHK H1K]) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1230 | show "Group.group GHK" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1231 | using GHK_def subgroups_Inter_pair subgroup_imp_group assms by blast | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1232 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1233 | next | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1234 | have H1K_incl:"subgroup H1K (G\<lparr>carrier:= (H\<inter>K)\<rparr>)" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1235 | proof(intro subgroup_incl) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1236 | show "subgroup H1K G" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1237 | using assms normal_imp_subgroup subgroups_Inter_pair incl_subgroup H1K_def by blast | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1238 | next | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1239 | show "subgroup (H\<inter>K) G" using HK_def subgroups_Inter_pair assms by auto | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1240 | next | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1241 | have "H1 \<subseteq> (carrier (G\<lparr>carrier:=H\<rparr>))" | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68445diff
changeset | 1242 | using assms(3) normal_imp_subgroup subgroup.subset by blast | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1243 | also have "... \<subseteq> H" by simp | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1244 | thus "H1K \<subseteq>H\<inter>K" | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1245 | using H1K_def calculation by auto | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1246 | qed | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1247 | thus "subgroup H1K GHK" using GHK_def by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1248 | next | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1249 | show "\<And> x h. x\<in>carrier GHK \<Longrightarrow> h\<in>H1K \<Longrightarrow> x \<otimes>\<^bsub>GHK\<^esub> h \<otimes>\<^bsub>GHK\<^esub> inv\<^bsub>GHK\<^esub> x\<in> H1K" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1250 | proof- | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1251 | have invHK: "\<lbrakk>y\<in>HK\<rbrakk> \<Longrightarrow> inv\<^bsub>GHK\<^esub> y = inv\<^bsub>GH\<^esub> y" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1252 | using m_inv_consistent assms HK_def GH_def GHK_def subgroups_Inter_pair by simp | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1253 | have multHK : "\<lbrakk>x\<in>HK;y\<in>HK\<rbrakk> \<Longrightarrow> x \<otimes>\<^bsub>(G\<lparr>carrier:=HK\<rparr>)\<^esub> y = x \<otimes> y" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1254 | using HK_def by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1255 | fix x assume p: "x\<in>carrier GHK" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1256 | fix h assume p2 : "h:H1K" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1257 | have "carrier(GHK)\<subseteq>HK" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1258 | using GHK_def HK_def by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1259 | hence xHK:"x\<in>HK" using p by auto | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1260 | hence invx:"inv\<^bsub>GHK\<^esub> x = inv\<^bsub>GH\<^esub> x" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1261 | using invHK assms GHK_def HK_def GH_def m_inv_consistent subgroups_Inter_pair by simp | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1262 | have "H1\<subseteq>carrier(GH)" | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68445diff
changeset | 1263 | using assms GH_def normal_imp_subgroup subgroup.subset by blast | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1264 | hence hHK:"h\<in>HK" | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1265 | using p2 H1K_def HK_def GH_def by auto | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1266 | hence xhx_egal : "x \<otimes>\<^bsub>GHK\<^esub> h \<otimes>\<^bsub>GHK\<^esub> inv\<^bsub>GHK\<^esub>x = x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1267 | using invx invHK multHK GHK_def GH_def by auto | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1268 | have xH:"x\<in>carrier(GH)" | 
| 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1269 | using xHK HK_def GH_def by auto | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1270 | have hH:"h\<in>carrier(GH)" | 
| 68555 
22d51874f37d
a few more lemmas from Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 1271 | using hHK HK_def GH_def by auto | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1272 | have "(\<forall>x\<in>carrier (GH). \<forall>h\<in>H1. x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> H1)" | 
| 68687 | 1273 | using assms GH_def normal.inv_op_closed2 by fastforce | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1274 | hence INCL_1 : "x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> H1" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1275 | using xH H1K_def p2 by blast | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1276 | have " x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> HK" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1277 | using assms HK_def subgroups_Inter_pair hHK xHK | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1278 | by (metis GH_def inf.cobounded1 subgroup_def subgroup_incl) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1279 | hence " x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> K" using HK_def by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1280 | hence " x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> H1K" using INCL_1 H1K_def by auto | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1281 | thus "x \<otimes>\<^bsub>GHK\<^esub> h \<otimes>\<^bsub>GHK\<^esub> inv\<^bsub>GHK\<^esub> x \<in> H1K" using xhx_egal by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1282 | qed | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1283 | qed | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1284 | qed | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1285 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1286 | lemma (in group) normal_inter_subgroup: | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1287 | assumes "subgroup H G" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1288 | and "N \<lhd> G" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1289 | shows "(N\<inter>H) \<lhd> (G\<lparr>carrier := H\<rparr>)" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1290 | proof - | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1291 | define K where "K = carrier G" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1292 | have "G\<lparr>carrier := K\<rparr> = G" using K_def by auto | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1293 | moreover have "subgroup K G" using K_def subgroup_self by blast | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1294 | moreover have "normal N (G \<lparr>carrier :=K\<rparr>)" using assms K_def by simp | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1295 | ultimately have "N \<inter> H \<lhd> G\<lparr>carrier := K \<inter> H\<rparr>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1296 | using normal_inter[of K H N] assms(1) by blast | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68445diff
changeset | 1297 | moreover have "K \<inter> H = H" using K_def assms subgroup.subset by blast | 
| 68687 | 1298 | ultimately show "normal (N\<inter>H) (G\<lparr>carrier := H\<rparr>)" | 
| 1299 | by auto | |
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1300 | qed | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 1301 | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 1302 | end |