| author | Thomas Lindae <thomas.lindae@in.tum.de> | 
| Thu, 18 Jul 2024 01:18:43 +0200 | |
| changeset 81080 | 4aa4bd946f96 | 
| parent 77406 | c2013f617a70 | 
| child 81438 | 95c9af7483b1 | 
| permissions | -rw-r--r-- | 
| 
77406
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1  | 
(* Title: The Second Isomorphism Theorem for Groups  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2  | 
Author: Jakob von Raumer, Karlsruhe Institute of Technology  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
3  | 
Maintainer: Jakob von Raumer <jakob.raumer@student.kit.edu>  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
6  | 
theory SndIsomorphismGrp  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
7  | 
imports Coset  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
8  | 
begin  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
10  | 
section \<open>The Second Isomorphism Theorem for Groups\<close>  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
12  | 
text \<open>This theory provides a proof of the second isomorphism theorems for groups.  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
13  | 
The theorems consist of several facts about normal subgroups.\<close>  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
15  | 
text \<open>The first lemma states that whenever we have a subgroup @{term S} and
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
16  | 
a normal subgroup @{term H} of a group @{term G}, their intersection is normal
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
17  | 
in @{term G}\<close>
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
19  | 
locale second_isomorphism_grp = normal +  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
20  | 
fixes S:: "'a set"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
21  | 
assumes subgrpS: "subgroup S G"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
23  | 
context second_isomorphism_grp  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
24  | 
begin  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
26  | 
interpretation groupS: group "G\<lparr>carrier := S\<rparr>"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
27  | 
using subgrpS  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
28  | 
by (metis subgroup_imp_group)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
30  | 
lemma normal_subgrp_intersection_normal:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
31  | 
shows "S \<inter> H \<lhd> (G\<lparr>carrier := S\<rparr>)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
32  | 
proof(auto simp: groupS.normal_inv_iff)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
33  | 
  from subgrpS is_subgroup have "\<And>x. x \<in> {S, H} \<Longrightarrow> subgroup x G" by auto
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
34  | 
  hence "subgroup (\<Inter> {S, H}) G" using subgroups_Inter by blast
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
35  | 
hence "subgroup (S \<inter> H) G" by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
36  | 
moreover have "S \<inter> H \<subseteq> S" by simp  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
37  | 
ultimately show "subgroup (S \<inter> H) (G\<lparr>carrier := S\<rparr>)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
38  | 
by (simp add: subgroup_incl subgrpS)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
39  | 
next  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
40  | 
fix g h  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
41  | 
assume g: "g \<in> S" and hH: "h \<in> H" and hS: "h \<in> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
42  | 
from g hH subgrpS show "g \<otimes> h \<otimes> inv\<^bsub>G\<lparr>carrier := S\<rparr>\<^esub> g \<in> H"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
43  | 
by (metis inv_op_closed2 subgroup.mem_carrier m_inv_consistent)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
44  | 
from g hS subgrpS show "g \<otimes> h \<otimes> inv\<^bsub>G\<lparr>carrier := S\<rparr>\<^esub> g \<in> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
45  | 
by (metis subgroup.m_closed subgroup.m_inv_closed m_inv_consistent)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
46  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
47  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
48  | 
lemma normal_set_mult_subgroup:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
49  | 
shows "subgroup (H <#> S) G"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
50  | 
proof(rule subgroupI)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
51  | 
show "H <#> S \<subseteq> carrier G"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
52  | 
by (metis setmult_subset_G subgroup.subset subgrpS subset)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
53  | 
next  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
54  | 
have "\<one> \<in> H" "\<one> \<in> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
55  | 
using is_subgroup subgrpS subgroup.one_closed by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
56  | 
hence "\<one> \<otimes> \<one> \<in> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
57  | 
unfolding set_mult_def by blast  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
58  | 
  thus "H <#> S \<noteq> {}" by auto
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
59  | 
next  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
60  | 
fix g  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
61  | 
assume g: "g \<in> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
62  | 
then obtain h s where h: "h \<in> H" and s: "s \<in> S" and ghs: "g = h \<otimes> s" unfolding set_mult_def  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
63  | 
by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
64  | 
hence "s \<in> carrier G" by (metis subgroup.mem_carrier subgrpS)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
65  | 
with h ghs obtain h' where h': "h' \<in> H" and "g = s \<otimes> h'"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
66  | 
using coset_eq unfolding r_coset_def l_coset_def by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
67  | 
with s have "inv g = (inv h') \<otimes> (inv s)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
68  | 
by (metis inv_mult_group mem_carrier subgroup.mem_carrier subgrpS)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
69  | 
moreover from h' s subgrpS have "inv h' \<in> H" "inv s \<in> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
70  | 
using subgroup.m_inv_closed m_inv_closed by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
71  | 
ultimately show "inv g \<in> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
72  | 
unfolding set_mult_def by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
73  | 
next  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
74  | 
fix g g'  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
75  | 
assume g: "g \<in> H <#> S" and h: "g' \<in> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
76  | 
then obtain h h' s s' where hh'ss': "h \<in> H" "h' \<in> H" "s \<in> S" "s' \<in> S" and "g = h \<otimes> s" and "g' = h' \<otimes> s'"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
77  | 
unfolding set_mult_def by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
78  | 
hence "g \<otimes> g' = (h \<otimes> s) \<otimes> (h' \<otimes> s')" by metis  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
79  | 
also from hh'ss' have inG: "h \<in> carrier G" "h' \<in> carrier G" "s \<in> carrier G" "s' \<in> carrier G"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
80  | 
using subgrpS mem_carrier subgroup.mem_carrier by force+  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
81  | 
hence "(h \<otimes> s) \<otimes> (h' \<otimes> s') = h \<otimes> (s \<otimes> h') \<otimes> s'"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
82  | 
using m_assoc by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
83  | 
also from hh'ss' inG obtain h'' where h'': "h'' \<in> H" and "s \<otimes> h' = h'' \<otimes> s"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
84  | 
using coset_eq unfolding r_coset_def l_coset_def  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
85  | 
by fastforce  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
86  | 
hence "h \<otimes> (s \<otimes> h') \<otimes> s' = h \<otimes> (h'' \<otimes> s) \<otimes> s'"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
87  | 
by simp  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
88  | 
also from h'' inG have "... = (h \<otimes> h'') \<otimes> (s \<otimes> s')"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
89  | 
using m_assoc mem_carrier by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
90  | 
finally have "g \<otimes> g' = h \<otimes> h'' \<otimes> (s \<otimes> s')".  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
91  | 
moreover have "... \<in> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
92  | 
unfolding set_mult_def using h'' hh'ss' subgrpS subgroup.m_closed by fastforce  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
93  | 
ultimately show "g \<otimes> g' \<in> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
94  | 
by simp  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
95  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
96  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
97  | 
lemma H_contained_in_set_mult:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
98  | 
shows "H \<subseteq> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
99  | 
proof  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
100  | 
fix x  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
101  | 
assume x: "x \<in> H"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
102  | 
have "x \<otimes> \<one> \<in> H <#> S" unfolding set_mult_def  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
103  | 
using second_isomorphism_grp.subgrpS second_isomorphism_grp_axioms subgroup.one_closed x by force  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
104  | 
with x show "x \<in> H <#> S" by (metis mem_carrier r_one)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
105  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
107  | 
lemma S_contained_in_set_mult:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
108  | 
shows "S \<subseteq> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
109  | 
proof  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
110  | 
fix s  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
111  | 
assume s: "s \<in> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
112  | 
then have "\<one> \<otimes> s \<in> H <#> S" unfolding set_mult_def by force  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
113  | 
with s show "s \<in> H <#> S" using subgrpS subgroup.mem_carrier l_one by force  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
114  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
115  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
116  | 
lemma normal_intersection_hom:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
117  | 
shows "group_hom (G\<lparr>carrier := S\<rparr>) ((G\<lparr>carrier := H <#> S\<rparr>) Mod H) (\<lambda>g. H #> g)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
118  | 
proof -  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
119  | 
have "group ((G\<lparr>carrier := H <#> S\<rparr>) Mod H)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
120  | 
by (simp add: H_contained_in_set_mult normal.factorgroup_is_group normal_axioms  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
121  | 
normal_restrict_supergroup normal_set_mult_subgroup)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
122  | 
moreover  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
123  | 
  { fix g
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
124  | 
assume g: "g \<in> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
125  | 
with g have "g \<in> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
126  | 
using S_contained_in_set_mult by blast  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
127  | 
hence "H #> g \<in> carrier ((G\<lparr>carrier := H <#> S\<rparr>) Mod H)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
128  | 
unfolding FactGroup_def RCOSETS_def r_coset_def by auto }  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
129  | 
moreover  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
130  | 
have "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> H #> x \<otimes> y = H #> x <#> (H #> y)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
131  | 
using normal.rcos_sum normal_axioms subgroup.mem_carrier subgrpS by fastforce  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
132  | 
ultimately show ?thesis  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
133  | 
by (auto simp: group_hom_def group_hom_axioms_def hom_def)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
134  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
135  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
136  | 
lemma normal_intersection_hom_kernel:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
137  | 
shows "kernel (G\<lparr>carrier := S\<rparr>) ((G\<lparr>carrier := H <#> S\<rparr>) Mod H) (\<lambda>g. H #> g) = H \<inter> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
138  | 
proof -  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
139  | 
have "kernel (G\<lparr>carrier := S\<rparr>) ((G\<lparr>carrier := H <#> S\<rparr>) Mod H) (\<lambda>g. H #> g)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
140  | 
      = {g \<in> S. H #> g = \<one>\<^bsub>(G\<lparr>carrier := H <#> S\<rparr>) Mod H\<^esub>}" 
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
141  | 
unfolding kernel_def by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
142  | 
  also have "... = {g \<in> S. H #> g = H}" 
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
143  | 
unfolding FactGroup_def by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
144  | 
  also have "... = {g \<in> S. g \<in> H}"
 | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
145  | 
by (meson coset_join1 is_group rcos_const subgroupE(1) subgroup_axioms subgrpS subset_eq)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
146  | 
also have "... = H \<inter> S" by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
147  | 
finally show ?thesis.  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
148  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
150  | 
lemma normal_intersection_hom_surj:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
151  | 
shows "(\<lambda>g. H #> g) ` carrier (G\<lparr>carrier := S\<rparr>) = carrier ((G\<lparr>carrier := H <#> S\<rparr>) Mod H)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
152  | 
proof auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
153  | 
fix g  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
154  | 
assume "g \<in> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
155  | 
hence "g \<in> H <#> S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
156  | 
using S_contained_in_set_mult by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
157  | 
thus "H #> g \<in> carrier ((G\<lparr>carrier := H <#> S\<rparr>) Mod H)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
158  | 
unfolding FactGroup_def RCOSETS_def r_coset_def by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
159  | 
next  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
160  | 
fix x  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
161  | 
assume "x \<in> carrier (G\<lparr>carrier := H <#> S\<rparr> Mod H)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
162  | 
then obtain h s where h: "h \<in> H" and s: "s \<in> S" and "x = H #> (h \<otimes> s)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
163  | 
unfolding FactGroup_def RCOSETS_def r_coset_def set_mult_def by auto  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
164  | 
hence "x = (H #> h) #> s"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
165  | 
by (metis h s coset_mult_assoc mem_carrier subgroup.mem_carrier subgrpS subset)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
166  | 
also have "... = H #> s"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
167  | 
by (metis h is_group rcos_const)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
168  | 
finally have "x = H #> s".  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
169  | 
with s show "x \<in> (#>) H ` S"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
170  | 
by simp  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
171  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
172  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
173  | 
text \<open>Finally we can prove the actual isomorphism theorem:\<close>  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
174  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
175  | 
theorem normal_intersection_quotient_isom:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
176  | 
shows "(\<lambda>X. the_elem ((\<lambda>g. H #> g) ` X)) \<in> iso ((G\<lparr>carrier := S\<rparr>) Mod (H \<inter> S)) (((G\<lparr>carrier := H <#> S\<rparr>)) Mod H)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
177  | 
using normal_intersection_hom_kernel[symmetric] normal_intersection_hom normal_intersection_hom_surj  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
178  | 
by (metis group_hom.FactGroup_iso_set)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
180  | 
end  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
181  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
182  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
183  | 
corollary (in group) normal_subgroup_set_mult_closed:  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
184  | 
assumes "M \<lhd> G" and "N \<lhd> G"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
185  | 
shows "M <#> N \<lhd> G"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
186  | 
proof (rule normalI)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
187  | 
from assms show "subgroup (M <#> N) G"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
188  | 
using second_isomorphism_grp.normal_set_mult_subgroup normal_imp_subgroup  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
189  | 
unfolding second_isomorphism_grp_def second_isomorphism_grp_axioms_def by force  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
190  | 
next  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
191  | 
show "\<forall>x\<in>carrier G. M <#> N #> x = x <# (M <#> N)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
192  | 
proof  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
193  | 
fix x  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
194  | 
assume x: "x \<in> carrier G"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
195  | 
have "M <#> N #> x = M <#> (N #> x)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
196  | 
by (metis assms normal_inv_iff setmult_rcos_assoc subgroup.subset x)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
197  | 
also have "\<dots> = M <#> (x <# N)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
198  | 
by (metis assms(2) normal.coset_eq x)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
199  | 
also have "\<dots> = (M #> x) <#> N"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
200  | 
by (metis assms normal_imp_subgroup rcos_assoc_lcos subgroup.subset x)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
201  | 
also have "\<dots> = x <# (M <#> N)"  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
202  | 
by (simp add: assms normal.coset_eq normal_imp_subgroup setmult_lcos_assoc subgroup.subset x)  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
203  | 
finally show "M <#> N #> x = x <# (M <#> N)" .  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
204  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
205  | 
qed  | 
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
206  | 
|
| 
 
c2013f617a70
Importation of basic group theory results, due to Jakob von Raumer from his AFP entry Jordan-Hölder Theorem
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
207  | 
end  |