src/HOL/Algebra/Coset.thy
author wenzelm
Mon, 20 May 2024 15:43:51 +0200
changeset 80182 29f2b8ff84f3
parent 80067 c40bdfc84640
child 80914 d97fdabd9e2b
permissions -rw-r--r--
proper support for "isabelle update -D DIR": avoid accidental exclusion of select_dirs (amending e5dafe9e120f);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
     1
(*  Title:      HOL/Algebra/Coset.thy
68582
b9b9e2985878 more standard headers;
wenzelm
parents: 68555
diff changeset
     2
    Authors:    Florian Kammueller, L C Paulson, Stephan Hohe
b9b9e2985878 more standard headers;
wenzelm
parents: 68555
diff changeset
     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: 68443
diff 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
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     7
theory Coset
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     8
imports Group
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     9
begin
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
    10
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
    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
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
    13
definition
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
    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: 35847
diff 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
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
    17
definition
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff 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: 64587
diff 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
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
    21
definition
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
    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: 35847
diff changeset
    23
  where "rcosets\<^bsub>G\<^esub> H = (\<Union>a\<in>carrier G. {H #>\<^bsub>G\<^esub> a})"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
    24
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
    25
definition
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff 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: 64587
diff 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
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
    29
definition
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
    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: 35847
diff 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
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
    33
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
    34
locale normal = subgroup + group +
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff 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
b808efaa5828 tuned syntax/abbreviations;
wenzelm
parents: 16417
diff changeset
    37
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20318
diff changeset
    38
  normal_rel :: "['a set, ('a, 'b) monoid_scheme] \<Rightarrow> bool"  (infixl "\<lhd>" 60) where
19380
b808efaa5828 tuned syntax/abbreviations;
wenzelm
parents: 16417
diff changeset
    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
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    41
lemma (in comm_group) subgroup_imp_normal: "subgroup A G \<Longrightarrow> A \<lhd> G"
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: 77362
diff changeset
    42
  by (simp add: normal_def normal_axioms_def l_coset_def r_coset_def m_comm subgroup.mem_carrier)
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    43
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
    44
lemma l_coset_eq_set_mult: \<^marker>\<open>contributor \<open>Martin Baillon\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    45
  fixes G (structure)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    46
  shows "x <# H = {x} <#> H"
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
    47
  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: 67443
diff changeset
    48
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
    49
lemma r_coset_eq_set_mult: \<^marker>\<open>contributor \<open>Martin Baillon\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    50
  fixes G (structure)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    51
  shows "H #> x = H <#> {x}"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    52
  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: 67443
diff changeset
    53
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
    54
lemma (in subgroup) rcosets_non_empty: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    55
  assumes "R \<in> rcosets H"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    56
  shows "R \<noteq> {}"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    57
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    58
  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: 67443
diff changeset
    59
    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: 67443
diff changeset
    60
  hence "\<one> \<otimes> g \<in> R"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    61
    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: 67443
diff changeset
    62
  thus ?thesis by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    63
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    64
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
    65
lemma (in group) diff_neutralizes: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    66
  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: 67443
diff changeset
    67
  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: 67443
diff changeset
    68
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    69
  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: 67443
diff changeset
    70
  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: 67443
diff changeset
    71
    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: 67443
diff changeset
    72
  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: 67443
diff changeset
    73
                      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: 67443
diff changeset
    74
    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: 67443
diff changeset
    75
  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: 68517
diff changeset
    76
    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: 67443
diff changeset
    77
  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: 67443
diff changeset
    78
    using h1 h2 assms(1) g(1) inv_closed m_closed monoid.m_assoc
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    79
          monoid_axioms subgroup.mem_carrier
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    80
  proof -
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    81
    have "h1 \<in> carrier G"
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    82
      by (meson subgroup.mem_carrier assms(1) h1(1))
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    83
    moreover have "h2 \<in> carrier G"
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    84
      by (meson subgroup.mem_carrier assms(1) h2(1))
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    85
    ultimately show ?thesis
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    86
      using g(1) inv_closed m_assoc m_closed by presburger
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
    87
  qed
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    88
  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: 67443
diff changeset
    89
    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: 67443
diff changeset
    90
  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: 67443
diff changeset
    91
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    92
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
    93
lemma mono_set_mult: "\<lbrakk> H \<subseteq> H'; K \<subseteq> K' \<rbrakk> \<Longrightarrow> H <#>\<^bsub>G\<^esub> K \<subseteq> H' <#>\<^bsub>G\<^esub> K'" \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
    94
  unfolding set_mult_def by (simp add: UN_mono)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
    95
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    96
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    97
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: 67443
diff changeset
    98
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
    99
lemma set_mult_consistent [simp]: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
   100
  "N <#>\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> K = N <#>\<^bsub>G\<^esub> K"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
   101
  unfolding set_mult_def by simp
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
   102
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
   103
lemma r_coset_consistent [simp]: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
   104
  "I #>\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> h = I #>\<^bsub>G\<^esub> h"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
   105
  unfolding r_coset_def by simp
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
   106
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
   107
lemma l_coset_consistent [simp]: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
   108
  "h <#\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> I = h <#\<^bsub>G\<^esub> I"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
   109
  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: 67443
diff changeset
   110
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff 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: 68443
diff 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: 67443
diff changeset
   113
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff 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: 67443
diff 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: 67443
diff 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: 67443
diff 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: 67443
diff changeset
   118
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff 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: 67443
diff 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: 67443
diff 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: 67443
diff 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: 67443
diff changeset
   123
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
   124
lemma (in group) set_mult_assoc: \<^marker>\<open>contributor \<open>Martin Baillon\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   125
  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: 67443
diff changeset
   126
  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: 67443
diff changeset
   127
proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   128
  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: 67443
diff changeset
   129
  proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   130
    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: 67443
diff changeset
   131
    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: 67443
diff changeset
   132
      unfolding set_mult_def by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   133
    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: 67443
diff changeset
   134
      using assms m_assoc by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   135
    thus "x \<in> M <#> (H <#> K)"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   136
      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: 67443
diff changeset
   137
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   138
next
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   139
  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: 67443
diff changeset
   140
  proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   141
    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: 67443
diff changeset
   142
    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: 67443
diff changeset
   143
      unfolding set_mult_def by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   144
    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: 67443
diff changeset
   145
      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: 67443
diff changeset
   146
    thus "x \<in> (M <#> H) <#> K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   147
      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: 67443
diff changeset
   148
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   149
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   150
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   151
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   152
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   153
subsection \<open>Basic Properties of Cosets\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   154
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   155
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: 67443
diff changeset
   156
  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: 67443
diff changeset
   157
  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: 67443
diff changeset
   158
  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: 67443
diff changeset
   159
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   160
lemma (in group) coset_assoc:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   161
  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: 67443
diff changeset
   162
  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: 67443
diff changeset
   163
  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: 67443
diff changeset
   164
  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
   165
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   166
lemma (in group) coset_mult_one [simp]: "M \<subseteq> carrier G ==> M #> \<one> = M"
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   167
by (force simp add: r_coset_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   168
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   169
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: 67443
diff changeset
   170
  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: 67443
diff changeset
   171
    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: 67443
diff changeset
   172
  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: 67443
diff changeset
   173
  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
   174
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   175
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: 67443
diff changeset
   176
  assumes "M #> x = M #> y"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   177
    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: 67443
diff changeset
   178
  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: 68517
diff changeset
   179
  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
   180
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   181
lemma (in group) coset_join1:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   182
  assumes "H #> x = H"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   183
    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: 67443
diff changeset
   184
  shows "x \<in> H"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   185
  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
   186
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   187
lemma (in group) solve_equation:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   188
  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: 67443
diff changeset
   189
  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: 67443
diff changeset
   190
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   191
  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: 67443
diff changeset
   192
    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: 67443
diff changeset
   193
  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: 67443
diff changeset
   194
    by (simp add: subgroup_def)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   195
  ultimately show ?thesis by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   196
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   197
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   198
lemma (in group_hom) inj_on_one_iff:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   199
   "inj_on h (carrier G) \<longleftrightarrow> (\<forall>x. x \<in> carrier G \<longrightarrow> h x = one H \<longrightarrow> x = one G)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   200
using G.solve_equation G.subgroup_self by (force simp: inj_on_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   201
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   202
lemma inj_on_one_iff':
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   203
   "\<lbrakk>h \<in> hom G H; group G; group H\<rbrakk> \<Longrightarrow> inj_on h (carrier G) \<longleftrightarrow> (\<forall>x. x \<in> carrier G \<longrightarrow> h x = one H \<longrightarrow> x = one G)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   204
  using group_hom.inj_on_one_iff group_hom.intro group_hom_axioms.intro by blast
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   205
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   206
lemma mon_iff_hom_one:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   207
   "\<lbrakk>group G; group H\<rbrakk> \<Longrightarrow> f \<in> mon G H \<longleftrightarrow> f \<in> hom G H \<and> (\<forall>x. x \<in> carrier G \<and> f x = \<one>\<^bsub>H\<^esub> \<longrightarrow> x = \<one>\<^bsub>G\<^esub>)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   208
  by (auto simp: mon_def inj_on_one_iff')
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   209
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   210
lemma (in group_hom) iso_iff: "h \<in> iso G H \<longleftrightarrow> carrier H \<subseteq> h ` carrier G \<and> (\<forall>x\<in>carrier G. h x = \<one>\<^bsub>H\<^esub> \<longrightarrow> x = \<one>)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   211
  by (auto simp: iso_def bij_betw_def inj_on_one_iff)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   212
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   213
lemma (in group) repr_independence:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   214
  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: 67443
diff changeset
   215
  shows "H #> x = H #> y" using assms
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   216
by (auto simp add: r_coset_def m_assoc [symmetric]
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   217
                   subgroup.subset [THEN subsetD]
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   218
                   subgroup.m_closed solve_equation)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   219
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   220
lemma (in group) coset_join2:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   221
  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: 67443
diff changeset
   222
  shows "H #> x = H" using assms
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
   223
  \<comment> \<open>Alternative proof is to put \<^term>\<open>x=\<one>\<close> in \<open>repr_independence\<close>.\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   224
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
   225
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   226
lemma (in group) coset_join3:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   227
  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: 67443
diff changeset
   228
  shows "x <# H = H"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   229
proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   230
  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: 67443
diff changeset
   231
    by (simp add: subgroup.m_closed)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   232
  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: 67443
diff changeset
   233
next
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   234
  have "\<And>h. h \<in> H \<Longrightarrow> x \<otimes> ((inv x) \<otimes> h) = h"
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   235
    by (metis (no_types, lifting) assms group.inv_closed group.inv_solve_left is_group 
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   236
              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: 67443
diff changeset
   237
  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: 67443
diff changeset
   238
    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: 67443
diff changeset
   239
  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: 67443
diff changeset
   240
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   241
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   242
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: 67443
diff changeset
   243
  "\<lbrakk> H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> H #> x \<subseteq> carrier G"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   244
by (auto simp add: r_coset_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   245
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   246
lemma (in group) rcosI:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   247
  "\<lbrakk> h \<in> H; H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> h \<otimes> x \<in> H #> x"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   248
by (auto simp add: r_coset_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   249
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   250
lemma (in group) rcosetsI:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   251
     "\<lbrakk>H \<subseteq> carrier G; x \<in> carrier G\<rbrakk> \<Longrightarrow> H #> x \<in> rcosets H"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   252
by (auto simp add: RCOSETS_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   253
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   254
lemma (in group) rcos_self:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   255
  "\<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: 67443
diff changeset
   256
  by (metis l_one rcosI subgroup_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   257
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   258
text (in group) \<open>Opposite of @{thm [source] "repr_independence"}\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   259
lemma (in group) repr_independenceD:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   260
  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: 67443
diff changeset
   261
    and "H #> x = H #> y"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   262
  shows "y \<in> H #> x"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   263
  using assms by (simp add: rcos_self)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   264
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   265
text \<open>Elements of a right coset are in the carrier\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   266
lemma (in subgroup) elemrcos_carrier:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   267
  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: 67443
diff changeset
   268
    and "a' \<in> H #> a"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   269
  shows "a' \<in> carrier G"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   270
  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: 19931
diff changeset
   271
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   272
lemma (in subgroup) rcos_const:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   273
  assumes "group G" "h \<in> H"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   274
  shows "H #> h = H"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   275
  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: 67443
diff changeset
   276
  by (simp add: assms(2) subgroup_axioms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   277
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   278
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: 67443
diff changeset
   279
  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: 67443
diff changeset
   280
    and "x' \<in> H #> x"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   281
  shows "(x' \<otimes> inv x) \<in> H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   282
proof -
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   283
  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: 67443
diff changeset
   284
    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: 67443
diff changeset
   285
  hence "x' \<otimes> inv x = h"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   286
    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: 67443
diff changeset
   287
  thus ?thesis using h by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   288
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   289
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   290
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: 67443
diff changeset
   291
  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: 67443
diff changeset
   292
    and "(x' \<otimes> inv x) \<in> H"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   293
  shows "x' \<in> H #> x"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   294
proof -
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   295
  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: 67443
diff changeset
   296
    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: 67443
diff changeset
   297
  hence "x' = h \<otimes> x"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   298
    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: 67443
diff changeset
   299
  thus ?thesis using h unfolding r_coset_def by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   300
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   301
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   302
text \<open>Module property of right cosets\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   303
lemma (in subgroup) rcos_module:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   304
  assumes "group G" "x \<in> carrier G" "x' \<in> carrier G"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   305
  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: 67443
diff changeset
   306
  using rcos_module_rev rcos_module_imp assms by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   307
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   308
text \<open>Right cosets are subsets of the carrier.\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   309
lemma (in subgroup) rcosets_carrier:
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   310
  assumes "group G" "X \<in> rcosets H"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   311
  shows "X \<subseteq> carrier G"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   312
  using assms elemrcos_carrier singletonD
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   313
  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: 67443
diff changeset
   314
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   315
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   316
text \<open>Multiplication of general subsets\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   317
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   318
lemma (in comm_group) mult_subgroups:
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   319
  assumes HG: "subgroup H G" and KG: "subgroup K G"
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   320
  shows "subgroup (H <#> K) G"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   321
proof (rule subgroup.intro)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   322
  show "H <#> K \<subseteq> carrier G"
68452
c027dfbfad30 more on infinite products. Also subgroup_imp_subset -> subgroup.subset
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
   323
    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: 67443
diff changeset
   324
next
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   325
  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: 67443
diff changeset
   326
    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: 67443
diff changeset
   327
  thus "\<one> \<in> H <#> K" by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   328
next
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   329
  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: 67443
diff changeset
   330
  proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   331
    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: 67443
diff changeset
   332
    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: 67443
diff changeset
   333
      unfolding set_mult_def by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   334
    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: 67443
diff changeset
   335
      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: 67443
diff changeset
   336
    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: 67443
diff changeset
   337
      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: 67443
diff changeset
   338
    thus "inv x \<in> H <#> K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   339
      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: 67443
diff changeset
   340
      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: 67443
diff changeset
   341
  qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   342
next
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   343
  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: 67443
diff changeset
   344
  proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   345
    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: 67443
diff changeset
   346
    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: 67443
diff changeset
   347
                              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: 67443
diff changeset
   348
      unfolding set_mult_def by blast
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   349
    with KG HG have carr: "k1 \<in> carrier G" "h1 \<in> carrier G" "k2 \<in> carrier G" "h2 \<in> carrier G"
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   350
        by (meson subgroup.mem_carrier)+
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   351
    have "x \<otimes> y = (h1 \<otimes> k1) \<otimes> (h2 \<otimes> k2)"
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   352
      using h1k1 h2k2 by simp
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   353
    also have " ... = h1 \<otimes> (k1 \<otimes> h2) \<otimes> k2"
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   354
        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: 67443
diff changeset
   355
    also have " ... = h1 \<otimes> (h2 \<otimes> k1) \<otimes> k2"
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   356
      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: 67443
diff changeset
   357
    finally have "x \<otimes> y  = (h1 \<otimes> h2) \<otimes> (k1 \<otimes> k2)"
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   358
      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: 67443
diff changeset
   359
    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: 67443
diff changeset
   360
      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: 67443
diff changeset
   361
            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: 67443
diff changeset
   362
  qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   363
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   364
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   365
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: 67443
diff changeset
   366
  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: 67443
diff changeset
   367
    and "(inv x \<otimes> x') \<in> H"
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   368
  shows "x' \<in> x <# H"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   369
proof -
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   370
  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: 67443
diff changeset
   371
    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: 67443
diff changeset
   372
  hence "x' = x \<otimes> h"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   373
    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: 67443
diff changeset
   374
  thus ?thesis using h unfolding l_coset_def by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   375
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   376
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   377
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   378
subsection \<open>Normal subgroups\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   379
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   380
lemma normal_imp_subgroup: "H \<lhd> G \<Longrightarrow> subgroup H G"
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   381
  by (rule normal.axioms(1))
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   382
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   383
lemma (in group) normalI:
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   384
  "subgroup H G \<Longrightarrow> (\<forall>x \<in> carrier G. H #> x = x <# H) \<Longrightarrow> H \<lhd> G"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   385
  by (simp add: normal_def normal_axioms_def is_group)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   386
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   387
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: 67443
diff changeset
   388
  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: 67443
diff changeset
   389
  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: 67443
diff changeset
   390
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   391
  have "h \<otimes> x \<in> x <# H"
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   392
    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: 67443
diff changeset
   393
  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: 67443
diff changeset
   394
    unfolding l_coset_def by blast
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   395
  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: 67443
diff changeset
   396
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   397
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   398
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: 67443
diff changeset
   399
  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: 67443
diff changeset
   400
  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: 68517
diff changeset
   401
  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: 67443
diff changeset
   402
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   403
lemma (in comm_group) normal_iff_subgroup:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   404
  "N \<lhd> G \<longleftrightarrow> subgroup N G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   405
proof
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   406
  assume "subgroup N G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   407
  then show "N \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   408
    by unfold_locales (auto simp: subgroupE subgroup.one_closed l_coset_def r_coset_def m_comm subgroup.mem_carrier)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   409
qed (simp add: normal_imp_subgroup)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   410
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   411
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   412
text\<open>Alternative characterization of normal subgroups\<close>
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   413
lemma (in group) normal_inv_iff:
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   414
     "(N \<lhd> G) =
67091
1393c2340eec more symbols;
wenzelm
parents: 65035
diff changeset
   415
      (subgroup N G \<and> (\<forall>x \<in> carrier G. \<forall>h \<in> N. x \<otimes> h \<otimes> (inv x) \<in> N))"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   416
      (is "_ = ?rhs")
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   417
proof
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   418
  assume N: "N \<lhd> G"
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   419
  show ?rhs
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   420
    by (blast intro: N normal.inv_op_closed2 normal_imp_subgroup)
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   421
next
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   422
  assume ?rhs
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   423
  hence sg: "subgroup N G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   424
    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: 68517
diff changeset
   425
  hence sb: "N \<subseteq> carrier G" by (simp add: subgroup.subset)
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   426
  show "N \<lhd> G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   427
  proof (intro normalI [OF sg], simp add: l_coset_def r_coset_def, clarify)
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   428
    fix x
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   429
    assume x: "x \<in> carrier G"
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   430
    show "(\<Union>h\<in>N. {h \<otimes> x}) = (\<Union>h\<in>N. {x \<otimes> h})"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   431
    proof
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   432
      show "(\<Union>h\<in>N. {h \<otimes> x}) \<subseteq> (\<Union>h\<in>N. {x \<otimes> h})"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   433
      proof clarify
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   434
        fix n
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   435
        assume n: "n \<in> N"
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   436
        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: 68517
diff changeset
   437
        proof
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   438
          from closed [of "inv x"]
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   439
          show "inv x \<otimes> n \<otimes> x \<in> N" by (simp add: x n)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   440
          show "n \<otimes> x \<in> {x \<otimes> (inv x \<otimes> n \<otimes> x)}"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   441
            by (simp add: x n m_assoc [symmetric] sb [THEN subsetD])
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   442
        qed
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   443
      qed
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   444
    next
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   445
      show "(\<Union>h\<in>N. {x \<otimes> h}) \<subseteq> (\<Union>h\<in>N. {h \<otimes> x})"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   446
      proof clarify
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   447
        fix n
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   448
        assume n: "n \<in> N"
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   449
        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: 68517
diff changeset
   450
        proof
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   451
          show "x \<otimes> n \<otimes> inv x \<in> N" by (simp add: x n closed)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   452
          show "x \<otimes> n \<in> {x \<otimes> n \<otimes> inv x \<otimes> x}"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   453
            by (simp add: x n m_assoc sb [THEN subsetD])
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   454
        qed
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   455
      qed
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   456
    qed
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   457
  qed
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   458
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   459
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   460
corollary (in group) normal_invI:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   461
  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: 67443
diff changeset
   462
  shows "N \<lhd> G"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   463
  using assms normal_inv_iff by blast
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   464
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
   465
corollary (in group) normal_invE:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
   466
  assumes "N \<lhd> G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
   467
  shows "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"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
   468
  using assms normal_inv_iff apply blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
   469
  by (simp add: assms normal.inv_op_closed2)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
   470
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: 77362
diff changeset
   471
lemma (in group) one_is_normal: "{\<one>} \<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: 77362
diff changeset
   472
  using normal_invI triv_subgroup by force
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
   473
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: 77362
diff changeset
   474
text \<open>The intersection of two normal subgroups is, again, a normal subgroup.\<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: 77362
diff changeset
   475
lemma (in group) normal_subgroup_intersect:
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: 77362
diff changeset
   476
  assumes "M \<lhd> G" and "N \<lhd> G" shows "M \<inter> 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: 77362
diff changeset
   477
  using  assms normal_inv_iff subgroups_Inter_pair 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: 77362
diff changeset
   478
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: 77362
diff changeset
   479
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: 77362
diff changeset
   480
text \<open>Being a normal subgroup is preserved by surjective homomorphisms.\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   481
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: 77362
diff changeset
   482
lemma (in normal) surj_hom_normal_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: 77362
diff changeset
   483
  assumes \<phi>: "group_hom G F \<phi>"
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: 77362
diff changeset
   484
  assumes \<phi>surj: "\<phi> ` (carrier G) = carrier F"
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: 77362
diff changeset
   485
  shows "(\<phi> ` H) \<lhd> F"
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: 77362
diff changeset
   486
proof (rule group.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: 77362
diff changeset
   487
  show "group F"
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: 77362
diff changeset
   488
    using \<phi> group_hom.axioms(2) 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: 77362
diff changeset
   489
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: 77362
diff changeset
   490
  show "subgroup (\<phi> ` H) F"
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: 77362
diff changeset
   491
    using \<phi> group_hom.subgroup_img_is_subgroup subgroup_axioms 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: 77362
diff changeset
   492
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: 77362
diff changeset
   493
  show "\<forall>x\<in>carrier F. \<phi> ` H #>\<^bsub>F\<^esub> x = x <#\<^bsub>F\<^esub> \<phi> ` 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: 77362
diff changeset
   494
  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: 77362
diff changeset
   495
    fix f
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: 77362
diff changeset
   496
    assume f: "f \<in> carrier F"
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: 77362
diff changeset
   497
    with \<phi>surj obtain g where g: "g \<in> carrier G" "f = \<phi> 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: 77362
diff changeset
   498
    hence "\<phi> ` H #>\<^bsub>F\<^esub> f = \<phi> ` H #>\<^bsub>F\<^esub> \<phi> g" 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: 77362
diff changeset
   499
    also have "... = (\<lambda>x. (\<phi> x) \<otimes>\<^bsub>F\<^esub> (\<phi> 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: 77362
diff changeset
   500
      unfolding r_coset_def image_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: 77362
diff changeset
   501
    also have "... = (\<lambda>x. \<phi> (x \<otimes> 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: 77362
diff changeset
   502
      using subset g \<phi> group_hom.hom_mult unfolding image_def 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: 77362
diff changeset
   503
    also have "... = \<phi> ` (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: 77362
diff changeset
   504
      using \<phi> unfolding 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: 77362
diff changeset
   505
    also have "... = \<phi> ` (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: 77362
diff changeset
   506
      by (metis coset_eq g(1))
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: 77362
diff changeset
   507
    also have "... = (\<lambda>x. \<phi> (g \<otimes> x)) ` 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: 77362
diff changeset
   508
      using \<phi> unfolding 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: 77362
diff changeset
   509
    also have "... = (\<lambda>x. (\<phi> g) \<otimes>\<^bsub>F\<^esub> (\<phi> x)) ` 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: 77362
diff changeset
   510
      using subset g \<phi> group_hom.hom_mult 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: 77362
diff changeset
   511
    also have "... = \<phi> g <#\<^bsub>F\<^esub> \<phi> ` 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: 77362
diff changeset
   512
      unfolding l_coset_def image_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: 77362
diff changeset
   513
    also have "... = f <#\<^bsub>F\<^esub> \<phi> ` 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: 77362
diff changeset
   514
      using g 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: 77362
diff changeset
   515
    finally show "\<phi> ` H #>\<^bsub>F\<^esub> f = f <#\<^bsub>F\<^esub> \<phi> ` 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: 77362
diff changeset
   516
  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: 77362
diff changeset
   517
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: 77362
diff changeset
   518
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: 77362
diff changeset
   519
text \<open>Being a normal subgroup is preserved by group isomorphisms.\<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: 77362
diff changeset
   520
lemma iso_normal_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: 77362
diff changeset
   521
  assumes \<phi>: "\<phi> \<in> iso G F" "group G" "group F" "H \<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: 77362
diff changeset
   522
  shows "(\<phi> ` H) \<lhd> F"
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: 77362
diff changeset
   523
  by (meson assms Group.iso_iff group_hom_axioms_def group_hom_def normal.surj_hom_normal_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: 77362
diff changeset
   524
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: 77362
diff changeset
   525
text \<open>The set product of two normal subgroups is a normal subgroup.\<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: 77362
diff changeset
   526
lemma (in group) setmult_lcos_assoc:
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: 77362
diff changeset
   527
  "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk>
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: 77362
diff changeset
   528
      \<Longrightarrow> (x <# H) <#> K = x <# (H <#> K)"
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: 77362
diff changeset
   529
  by (force simp add: l_coset_def set_mult_def m_assoc)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   530
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   531
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: 67443
diff changeset
   532
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   533
lemma (in group) l_repr_independence:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   534
  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: 67443
diff changeset
   535
  shows "x <# H = y <# H"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   536
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   537
  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: 67443
diff changeset
   538
    using assms(1) unfolding l_coset_def by blast
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   539
  hence "x \<otimes> h = y \<otimes> ((inv h') \<otimes> h)" if "h \<in> H" for h
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   540
  proof -
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   541
    have "h' \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   542
      by (meson HG h'(1) subgroup.mem_carrier)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   543
    moreover have "h \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   544
      by (meson HG subgroup.mem_carrier that)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   545
    ultimately show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   546
      by (metis assms(2) h'(2) inv_closed inv_solve_right m_assoc m_closed)
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   547
  qed
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   548
  hence "\<And>xh. xh \<in> x <# H \<Longrightarrow> xh \<in> y <# H"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   549
    unfolding l_coset_def by (metis (no_types, lifting) UN_iff HG h'(1) subgroup_def)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   550
  moreover have "\<And>h. h \<in> H \<Longrightarrow> y \<otimes> h = x \<otimes> (h' \<otimes> h)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   551
    using h' by (meson assms(2) HG m_assoc subgroup.mem_carrier)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   552
  hence "\<And>yh. yh \<in> y <# H \<Longrightarrow> yh \<in> x <# H"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   553
    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: 67443
diff changeset
   554
  ultimately show ?thesis by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   555
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   556
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   557
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: 67443
diff changeset
   558
  "\<lbrakk> M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G \<rbrakk> \<Longrightarrow> g <# (h <# M) = (g \<otimes> h) <# M"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   559
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
   560
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   561
lemma (in group) lcos_mult_one: "M \<subseteq> carrier G \<Longrightarrow> \<one> <# M = M"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   562
by (force simp add: l_coset_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   563
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   564
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: 67443
diff changeset
   565
  "\<lbrakk> H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> x <# H \<subseteq> carrier G"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   566
by (auto simp add: l_coset_def subsetD)
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   567
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   568
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: 67443
diff changeset
   569
  "\<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: 67443
diff changeset
   570
  by (auto simp add: l_coset_def m_assoc  subgroup.subset [THEN subsetD] subgroup.m_closed)
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   571
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   572
lemma (in group) l_coset_swap:
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   573
  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: 67443
diff changeset
   574
  shows "x \<in> y <# H"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   575
  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: 67443
diff changeset
   576
  unfolding l_coset_def by fastforce
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   577
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   578
lemma (in group) subgroup_mult_id:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   579
  assumes "subgroup H G"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   580
  shows "H <#> H = H"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   581
proof
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   582
  show "H <#> H \<subseteq> H"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   583
    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: 67443
diff changeset
   584
  show "H \<subseteq> H <#> H"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   585
  proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   586
    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: 67443
diff changeset
   587
      using subgroup.m_closed[OF assms subgroup.one_closed[OF assms] x] subgroup.one_closed[OF assms]
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   588
      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: 67443
diff changeset
   589
  qed
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   590
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   591
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   592
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   593
subsubsection \<open>Set of Inverses of an \<open>r_coset\<close>.\<close>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   594
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   595
lemma (in normal) rcos_inv:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   596
  assumes x:     "x \<in> carrier G"
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   597
  shows "set_inv (H #> x) = H #> (inv x)"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   598
proof (simp add: r_coset_def SET_INV_def x inv_mult_group, safe)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   599
  fix h
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   600
  assume h: "h \<in> H"
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   601
  show "inv x \<otimes> inv h \<in> (\<Union>j\<in>H. {j \<otimes> inv x})"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   602
  proof
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   603
    show "inv x \<otimes> inv h \<otimes> x \<in> H"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   604
      by (simp add: inv_op_closed1 h x)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   605
    show "inv x \<otimes> inv h \<in> {inv x \<otimes> inv h \<otimes> x \<otimes> inv x}"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   606
      by (simp add: h x m_assoc)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   607
  qed
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   608
  show "h \<otimes> inv x \<in> (\<Union>j\<in>H. {inv x \<otimes> inv j})"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   609
  proof
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   610
    show "x \<otimes> inv h \<otimes> inv x \<in> H"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   611
      by (simp add: inv_op_closed2 h x)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   612
    show "h \<otimes> inv x \<in> {inv x \<otimes> inv (x \<otimes> inv h \<otimes> inv x)}"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   613
      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
   614
  qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   615
qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   616
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   617
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   618
subsubsection \<open>Theorems for \<open><#>\<close> with \<open>#>\<close> or \<open><#\<close>.\<close>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   619
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   620
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: 67443
diff changeset
   621
  "\<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: 67443
diff changeset
   622
    H <#> (K #> x) = (H <#> K) #> x"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   623
  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
   624
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   625
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: 67443
diff changeset
   626
  "\<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: 67443
diff changeset
   627
   (H #> x) <#> K = H <#> (x <# K)"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   628
  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: 67443
diff changeset
   629
  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
   630
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   631
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: 67443
diff changeset
   632
  "\<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: 67443
diff changeset
   633
   (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: 67443
diff changeset
   634
  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: 67443
diff changeset
   635
                subset l_coset_subset_G rcos_assoc_lcos)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   636
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   637
lemma (in normal) rcos_mult_step2:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   638
     "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   639
      \<Longrightarrow> (H <#> (x <# H)) #> y = (H <#> (H #> x)) #> y"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   640
by (insert coset_eq, simp add: normal_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   641
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   642
lemma (in normal) rcos_mult_step3:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   643
     "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   644
      \<Longrightarrow> (H <#> (H #> x)) #> y = H #> (x \<otimes> y)"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   645
by (simp add: setmult_rcos_assoc coset_mult_assoc
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   646
              subgroup_mult_id normal.axioms subset normal_axioms)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   647
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   648
lemma (in normal) rcos_sum:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   649
     "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   650
      \<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
   651
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
   652
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   653
lemma (in normal) rcosets_mult_eq: "M \<in> rcosets H \<Longrightarrow> H <#> M = M"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   654
  \<comment> \<open>generalizes \<open>subgroup_mult_id\<close>\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   655
  by (auto simp add: RCOSETS_def subset
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   656
        setmult_rcos_assoc subgroup_mult_id normal.axioms normal_axioms)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   657
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   658
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   659
subsubsection\<open>An Equivalence Relation\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   660
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
   661
definition
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
   662
  r_congruent :: "[('a,'b)monoid_scheme, 'a set] \<Rightarrow> ('a*'a)set"  ("rcong\<index> _")
67091
1393c2340eec more symbols;
wenzelm
parents: 65035
diff changeset
   663
  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
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   664
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   665
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   666
lemma (in subgroup) equiv_rcong:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   667
   assumes "group G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   668
   shows "equiv (carrier G) (rcong H)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   669
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   670
  interpret group G by fact
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   671
  show ?thesis
40815
6e2d17cc0d1d equivI has replaced equiv.intro
haftmann
parents: 39910
diff changeset
   672
  proof (intro equivI)
80067
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 77407
diff changeset
   673
    have "rcong H \<subseteq> carrier G \<times> carrier G"
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 77407
diff changeset
   674
      by (auto simp add: r_congruent_def)
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 77407
diff changeset
   675
    thus "refl_on (carrier G) (rcong H)"
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   676
      by (auto simp add: r_congruent_def refl_on_def)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   677
  next
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   678
    show "sym (rcong H)"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   679
    proof (simp add: r_congruent_def sym_def, clarify)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   680
      fix x y
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   681
      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: 31727
diff changeset
   682
         and "inv x \<otimes> y \<in> H"
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 41528
diff changeset
   683
      hence "inv (inv x \<otimes> y) \<in> H" by simp
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   684
      thus "inv y \<otimes> x \<in> H" by (simp add: inv_mult_group)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   685
    qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   686
  next
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   687
    show "trans (rcong H)"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   688
    proof (simp add: r_congruent_def trans_def, clarify)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   689
      fix x y z
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   690
      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: 31727
diff changeset
   691
         and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   692
      hence "(inv x \<otimes> y) \<otimes> (inv y \<otimes> z) \<in> H" by simp
27698
197f0517f0bd Unit_inv_l, Unit_inv_r made [simp].
ballarin
parents: 27611
diff changeset
   693
      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: 68517
diff changeset
   694
        by (simp add: m_assoc del: r_inv Units_r_inv)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   695
      thus "inv x \<otimes> z \<in> H" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   696
    qed
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   697
  qed
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   698
qed
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   699
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   700
text\<open>Equivalence classes of \<open>rcong\<close> correspond to left cosets.
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   701
  Was there a mistake in the definitions? I'd have expected them to
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   702
  correspond to right cosets.\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   703
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   704
(* CB: This is correct, but subtle.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   705
   We call H #> a the right coset of a relative to H.  According to
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   706
   Jacobson, this is what the majority of group theory literature does.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   707
   He then defines the notion of congruence relation ~ over monoids as
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   708
   equivalence relation with a ~ a' & b ~ b' \<Longrightarrow> a*b ~ a'*b'.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   709
   Our notion of right congruence induced by K: rcong K appears only in
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   710
   the context where K is a normal subgroup.  Jacobson doesn't name it.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   711
   But in this context left and right cosets are identical.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   712
*)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   713
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   714
lemma (in subgroup) l_coset_eq_rcong:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   715
  assumes "group G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   716
  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: 67443
diff changeset
   717
  shows "a <# H = (rcong H) `` {a}"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   718
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   719
  interpret group G by fact
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   720
  show ?thesis by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a )
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   721
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   722
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   723
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   724
subsubsection\<open>Two Distinct Right Cosets are Disjoint\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   725
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   726
lemma (in group) rcos_equation:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   727
  assumes "subgroup H G"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   728
  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"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   729
  shows "hb \<otimes> a \<in> (\<Union>h\<in>H. {h \<otimes> b})"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   730
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   731
  interpret subgroup H G by fact
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   732
  from p show ?thesis 
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   733
    by (rule_tac UN_I [of "hb \<otimes> ((inv ha) \<otimes> h)"]) (auto simp: inv_solve_left m_assoc)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   734
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   735
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   736
lemma (in group) rcos_disjoint:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   737
  assumes "subgroup H G"
68975
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   738
  shows "pairwise disjnt (rcosets H)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   739
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   740
  interpret subgroup H G by fact
68975
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   741
  show ?thesis
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   742
    unfolding RCOSETS_def r_coset_def pairwise_def disjnt_def
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   743
    by (blast intro: rcos_equation assms sym)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   744
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   745
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   746
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   747
subsection \<open>Further lemmas for \<open>r_congruent\<close>\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   748
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   749
text \<open>The relation is a congruence\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   750
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   751
lemma (in normal) congruent_rcong:
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   752
  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: 19931
diff changeset
   753
proof (intro congruent2I[of "carrier G" _ "carrier G" _] equiv_rcong is_group)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   754
  fix a b c
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   755
  assume abrcong: "(a, b) \<in> rcong H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   756
    and ccarr: "c \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   757
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   758
  from abrcong
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   759
      have acarr: "a \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   760
        and bcarr: "b \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   761
        and abH: "inv a \<otimes> b \<in> H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   762
      unfolding r_congruent_def
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   763
      by fast+
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   764
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   765
  note carr = acarr bcarr ccarr
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   766
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   767
  from ccarr and abH
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   768
      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: 19931
diff changeset
   769
  moreover
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   770
      from carr and inv_closed
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   771
      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: 19931
diff changeset
   772
      by (force cong: m_assoc)
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   773
  moreover
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   774
      from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   775
      have "\<dots> = (inv (a \<otimes> c)) \<otimes> (b \<otimes> c)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   776
      by (simp add: inv_mult_group)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   777
  ultimately
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   778
      have "(inv (a \<otimes> c)) \<otimes> (b \<otimes> c) \<in> H" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   779
  from carr and this
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   780
     have "(b \<otimes> c) \<in> (a \<otimes> c) <# H"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   781
     by (simp add: lcos_module_rev[OF is_group])
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   782
  from carr and this and is_subgroup
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   783
     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: 19931
diff changeset
   784
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   785
  fix a b c
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   786
  assume abrcong: "(a, b) \<in> rcong H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   787
    and ccarr: "c \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   788
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 41528
diff changeset
   789
  from ccarr have "c \<in> Units G" by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   790
  hence cinvc_one: "inv c \<otimes> c = \<one>" by (rule Units_l_inv)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   791
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   792
  from abrcong
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   793
      have acarr: "a \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   794
       and bcarr: "b \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   795
       and abH: "inv a \<otimes> b \<in> H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   796
      by (unfold r_congruent_def, fast+)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   797
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   798
  note carr = acarr bcarr ccarr
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   799
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   800
  from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   801
     have "inv a \<otimes> b = inv a \<otimes> (\<one> \<otimes> b)" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   802
  also from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   803
      have "\<dots> = inv a \<otimes> (inv c \<otimes> c) \<otimes> b" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   804
  also from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   805
      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: 19931
diff changeset
   806
  also from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   807
      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: 19931
diff changeset
   808
  finally
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   809
      have "inv a \<otimes> b = inv (c \<otimes> a) \<otimes> (c \<otimes> b)" .
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   810
  from abH and this
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   811
      have "inv (c \<otimes> a) \<otimes> (c \<otimes> b) \<in> H" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   812
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   813
  from carr and this
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   814
     have "(c \<otimes> b) \<in> (c \<otimes> a) <# H"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   815
     by (simp add: lcos_module_rev[OF is_group])
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   816
  from carr and this and is_subgroup
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   817
     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: 19931
diff changeset
   818
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   819
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   820
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   821
subsection \<open>Order of a Group and Lagrange's Theorem\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   822
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   823
definition
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   824
  order :: "('a, 'b) monoid_scheme \<Rightarrow> nat"
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   825
  where "order S = card (carrier S)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   826
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: 77362
diff changeset
   827
lemma iso_same_order:
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: 77362
diff changeset
   828
  assumes "\<phi> \<in> iso 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: 77362
diff changeset
   829
  shows "order G = order 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: 77362
diff changeset
   830
  by (metis assms is_isoI iso_same_card order_def order_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: 77362
diff changeset
   831
61628
8dd2bd4fe30b add lemmas about monoids and groups
Andreas Lochbihler
parents: 61382
diff changeset
   832
lemma (in monoid) order_gt_0_iff_finite: "0 < order G \<longleftrightarrow> finite (carrier G)"
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: 77362
diff changeset
   833
  by(auto simp add: order_def card_gt_0_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: 77362
diff changeset
   834
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: 77362
diff changeset
   835
lemma (in group) order_one_triv_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: 77362
diff changeset
   836
  shows "(order G = 1) = (carrier G = {\<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: 77362
diff changeset
   837
  by (metis One_nat_def card.empty card_Suc_eq empty_iff one_closed order_def singleton_iff)
61628
8dd2bd4fe30b add lemmas about monoids and groups
Andreas Lochbihler
parents: 61382
diff changeset
   838
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   839
lemma (in group) rcosets_part_G:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   840
  assumes "subgroup H G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   841
  shows "\<Union>(rcosets H) = carrier G"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   842
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   843
  interpret subgroup H G by fact
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   844
  show ?thesis
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   845
    unfolding RCOSETS_def r_coset_def by auto
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   846
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   847
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   848
lemma (in group) cosets_finite:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   849
     "\<lbrakk>c \<in> rcosets H;  H \<subseteq> carrier G;  finite (carrier G)\<rbrakk> \<Longrightarrow> finite c"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   850
  unfolding RCOSETS_def
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   851
  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
   852
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   853
text\<open>The next two lemmas support the proof of \<open>card_cosets_equal\<close>.\<close>
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   854
lemma (in group) inj_on_f:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   855
  assumes "H \<subseteq> carrier G" and a: "a \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   856
  shows "inj_on (\<lambda>y. y \<otimes> inv a) (H #> a)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   857
proof 
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   858
  fix x y
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   859
  assume "x \<in> H #> a" "y \<in> H #> a" and xy: "x \<otimes> inv a = y \<otimes> inv a"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   860
  then have "x \<in> carrier G" "y \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   861
    using assms r_coset_subset_G by blast+
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   862
  with xy a show "x = y"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   863
    by auto
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   864
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   865
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   866
lemma (in group) inj_on_g:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   867
    "\<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
   868
by (force simp add: inj_on_def subsetD)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   869
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   870
(* ************************************************************************** *)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   871
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   872
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: 67443
diff changeset
   873
  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: 67443
diff changeset
   874
  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: 67443
diff changeset
   875
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   876
  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: 67443
diff changeset
   877
    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: 67443
diff changeset
   878
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   879
  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: 67443
diff changeset
   880
  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: 67443
diff changeset
   881
  proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   882
    fix r assume "r \<in> R"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   883
    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: 67443
diff changeset
   884
      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: 67443
diff changeset
   885
    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: 67443
diff changeset
   886
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   887
  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: 67443
diff changeset
   888
  moreover have "?f ` H \<subseteq> R"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   889
    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: 67443
diff changeset
   890
  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: 68517
diff changeset
   891
    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: 67443
diff changeset
   892
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   893
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   894
corollary (in group) card_rcosets_equal:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   895
  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: 67443
diff changeset
   896
  shows "card H = card R"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   897
  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: 67443
diff changeset
   898
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   899
corollary (in group) rcosets_finite:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   900
  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: 67443
diff changeset
   901
  shows "finite R"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   902
  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: 67443
diff changeset
   903
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   904
(* ************************************************************************** *)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   905
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   906
lemma (in group) rcosets_subset_PowG:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   907
     "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: 67443
diff changeset
   908
  using rcosets_part_G by auto
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   909
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   910
proposition (in group) lagrange_finite:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   911
  assumes "finite(carrier G)" and HG: "subgroup H G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   912
  shows "card(rcosets H) * card(H) = order(G)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   913
proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   914
  have "card H * card (rcosets H) = card (\<Union>(rcosets H))"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   915
  proof (rule card_partition)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   916
    show "\<And>c1 c2. \<lbrakk>c1 \<in> rcosets H; c2 \<in> rcosets H; c1 \<noteq> c2\<rbrakk> \<Longrightarrow> c1 \<inter> c2 = {}"
68975
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   917
      using HG rcos_disjoint by (auto simp: pairwise_def disjnt_def)
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   918
  qed (auto simp: assms finite_UnionD rcosets_part_G card_rcosets_equal subgroup.subset)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   919
  then show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   920
    by (simp add: HG mult.commute order_def rcosets_part_G)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   921
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   922
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   923
theorem (in group) lagrange:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   924
  assumes "subgroup H G"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   925
  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: 67443
diff changeset
   926
proof (cases "finite (carrier G)")
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   927
  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: 67443
diff changeset
   928
next
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   929
  case False 
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   930
  thus ?thesis
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   931
  proof (cases "finite H")
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   932
    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: 67443
diff changeset
   933
  next
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   934
    case True 
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   935
    have "infinite (rcosets H)"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   936
    proof 
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   937
      assume "finite (rcosets H)"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   938
      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: 67443
diff changeset
   939
      hence "card (\<Union>(rcosets H)) = (\<Sum>R\<in>(rcosets H). card R)"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   940
        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: 68445
diff changeset
   941
              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: 67443
diff changeset
   942
      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: 67443
diff changeset
   943
        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: 67443
diff changeset
   944
      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: 68445
diff changeset
   945
        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: 67443
diff changeset
   946
      hence "order G = (card H) * (card (rcosets H))" by simp
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   947
      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: 67443
diff changeset
   948
                                rcosets_part_G subgroup.one_closed by fastforce
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   949
      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: 67443
diff changeset
   950
    qed
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   951
    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: 67443
diff changeset
   952
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   953
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   954
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: 77362
diff changeset
   955
text \<open>The cardinality of the right cosets of the trivial subgroup is the cardinality of the group itself:\<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: 77362
diff changeset
   956
corollary (in group) card_rcosets_triv:
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: 77362
diff changeset
   957
  assumes "finite (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: 77362
diff changeset
   958
  shows "card (rcosets {\<one>}) = order 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: 77362
diff changeset
   959
  using lagrange triv_subgroup by fastforce
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   960
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   961
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
   962
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   963
definition
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   964
  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: 67091
diff changeset
   965
    \<comment> \<open>Actually defined for groups rather than monoids\<close>
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   966
   where "FactGroup G H = \<lparr>carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   967
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   968
lemma (in normal) setmult_closed:
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   969
     "\<lbrakk>K1 \<in> rcosets H; K2 \<in> rcosets H\<rbrakk> \<Longrightarrow> K1 <#> K2 \<in> rcosets H"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   970
by (auto simp add: rcos_sum RCOSETS_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   971
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   972
lemma (in normal) setinv_closed:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   973
     "K \<in> rcosets H \<Longrightarrow> set_inv K \<in> rcosets H"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   974
by (auto simp add: rcos_inv RCOSETS_def)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   975
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   976
lemma (in normal) rcosets_assoc:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   977
     "\<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: 64587
diff changeset
   978
      \<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: 67443
diff changeset
   979
  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
   980
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   981
lemma (in subgroup) subgroup_in_rcosets:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   982
  assumes "group G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   983
  shows "H \<in> rcosets H"
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   984
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   985
  interpret group G by fact
26203
9625f3579b48 explicit referencing of background facts;
wenzelm
parents: 23463
diff changeset
   986
  from _ subgroup_axioms have "H #> \<one> = H"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   987
    by (rule coset_join2) auto
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   988
  then show ?thesis
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   989
    by (auto simp add: RCOSETS_def)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   990
qed
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   991
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   992
lemma (in normal) rcosets_inv_mult_group_eq:
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   993
     "M \<in> rcosets H \<Longrightarrow> set_inv M <#> M = H"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   994
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: 13870
diff changeset
   995
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   996
theorem (in normal) factorgroup_is_group: "group (G Mod H)"
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   997
proof -
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   998
  have "\<And>x. x \<in> rcosets H \<Longrightarrow> \<exists>y\<in>rcosets H. y <#> x = H"
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   999
    using rcosets_inv_mult_group_eq setinv_closed by blast
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1000
  then show ?thesis
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1001
    unfolding FactGroup_def
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1002
    by (intro groupI)
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1003
      (auto simp: setmult_closed subgroup_in_rcosets rcosets_assoc rcosets_mult_eq)
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1004
qed
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
  1005
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1006
lemma carrier_FactGroup: "carrier(G Mod N) = (\<lambda>x. r_coset G N x) ` carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1007
  by (auto simp: FactGroup_def RCOSETS_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1008
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1009
lemma one_FactGroup [simp]: "one(G Mod N) = N"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1010
  by (auto simp: FactGroup_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1011
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1012
lemma mult_FactGroup [simp]: "monoid.mult (G Mod N) = set_mult G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1013
  by (auto simp: FactGroup_def)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1014
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1015
lemma (in normal) inv_FactGroup:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1016
  assumes "X \<in> carrier (G Mod H)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1017
  shows "inv\<^bsub>G Mod H\<^esub> X = set_inv X"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1018
proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1019
  have X: "X \<in> rcosets H"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1020
    using assms by (simp add: FactGroup_def)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1021
  moreover have "set_inv X <#> X = H"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1022
    using X by (simp add: normal.rcosets_inv_mult_group_eq normal_axioms)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1023
  moreover have "Group.group (G Mod H)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1024
    using normal.factorgroup_is_group normal_axioms by blast
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1025
  ultimately show ?thesis
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1026
    by (simp add: FactGroup_def group.inv_equality normal.setinv_closed normal_axioms)
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1027
qed
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
  1028
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
  1029
text\<open>The coset map is a homomorphism from \<^term>\<open>G\<close> to the quotient group
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
  1030
  \<^term>\<open>G Mod H\<close>\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1031
lemma (in normal) r_coset_hom_Mod:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1032
  "(\<lambda>a. H #> a) \<in> hom G (G Mod H)"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1033
  by (auto simp add: FactGroup_def RCOSETS_def Pi_def hom_def rcos_sum)
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
  1034
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1035
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1036
lemma (in comm_group) set_mult_commute:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1037
  assumes "N \<subseteq> carrier G" "x \<in> rcosets N" "y \<in> rcosets N"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1038
  shows "x <#> y = y <#> x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1039
  using assms unfolding set_mult_def RCOSETS_def
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1040
  by auto (metis m_comm r_coset_subset_G subsetCE)+
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1041
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1042
lemma (in comm_group) abelian_FactGroup:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1043
  assumes "subgroup N G" shows "comm_group(G Mod N)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1044
proof (rule group.group_comm_groupI)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1045
  have "N \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1046
    by (simp add: assms normal_iff_subgroup)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1047
  then show "Group.group (G Mod N)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1048
    by (simp add: normal.factorgroup_is_group)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1049
  fix x :: "'a set" and y :: "'a set"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1050
  assume "x \<in> carrier (G Mod N)" "y \<in> carrier (G Mod N)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1051
  then show "x \<otimes>\<^bsub>G Mod N\<^esub> y = y \<otimes>\<^bsub>G Mod N\<^esub> x"
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1052
    by (metis FactGroup_def assms mult_FactGroup partial_object.simps(1) set_mult_commute subgroup_def)
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1053
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1054
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1055
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1056
lemma FactGroup_universal:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1057
  assumes "h \<in> hom G H" "N \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1058
    and h: "\<And>x y. \<lbrakk>x \<in> carrier G; y \<in> carrier G; r_coset G N x = r_coset G N y\<rbrakk> \<Longrightarrow> h x = h y"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1059
  obtains g
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1060
  where "g \<in> hom (G Mod N) H" "\<And>x. x \<in> carrier G \<Longrightarrow> g(r_coset G N x) = h x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1061
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1062
  obtain g where g: "\<And>x. x \<in> carrier G \<Longrightarrow> h x = g(r_coset G N x)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1063
    using h function_factors_left_gen [of "\<lambda>x. x \<in> carrier G" "r_coset G N" h] by blast
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1064
  show thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1065
  proof
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1066
    show "g \<in> hom (G Mod N) H"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1067
    proof (rule homI)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1068
      show "g (u \<otimes>\<^bsub>G Mod N\<^esub> v) = g u \<otimes>\<^bsub>H\<^esub> g v"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1069
        if "u \<in> carrier (G Mod N)" "v \<in> carrier (G Mod N)" for u v
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1070
      proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1071
        from that
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1072
        obtain x y where xy: "x \<in> carrier G" "u = r_coset G N x" "y \<in> carrier G"  "v = r_coset G N y"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1073
          by (auto simp: carrier_FactGroup)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1074
        then have "h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1075
           by (metis hom_mult [OF \<open>h \<in> hom G H\<close>])
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1076
        then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1077
          by (metis Coset.mult_FactGroup xy \<open>N \<lhd> G\<close> g group.subgroup_self normal.axioms(2) normal.rcos_sum subgroup_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1078
      qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1079
    qed (use \<open>h \<in> hom G H\<close> in \<open>auto simp: carrier_FactGroup Pi_iff hom_def simp flip: g\<close>)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1080
  qed (auto simp flip: g)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1081
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1082
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1083
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1084
lemma (in normal) FactGroup_pow:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1085
  fixes k::nat
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1086
  assumes "a \<in> carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1087
  shows "pow (FactGroup G H) (r_coset G H a) k = r_coset G H (pow G a k)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1088
proof (induction k)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1089
  case 0
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1090
  then show ?case
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1091
    by (simp add: r_coset_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1092
next
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1093
  case (Suc k)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1094
  then show ?case
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1095
    by (simp add: assms rcos_sum)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1096
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1097
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1098
lemma (in normal) FactGroup_int_pow:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1099
  fixes k::int
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1100
  assumes "a \<in> carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1101
  shows "pow (FactGroup G H) (r_coset G H a) k = r_coset G H (pow G a k)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1102
  by (metis Group.group.axioms(1) image_eqI is_group monoid.nat_pow_closed int_pow_def2 assms
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1103
         FactGroup_pow carrier_FactGroup inv_FactGroup rcos_inv)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1104
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1105
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1106
subsection\<open>The First Isomorphism Theorem\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1107
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1108
text\<open>The quotient by the kernel of a homomorphism is isomorphic to the
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1109
  range of that homomorphism.\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1110
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
  1111
definition
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
  1112
  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: 67091
diff changeset
  1113
    \<comment> \<open>the kernel of a homomorphism\<close>
67091
1393c2340eec more symbols;
wenzelm
parents: 65035
diff changeset
  1114
  where "kernel G H h = {x. x \<in> carrier G \<and> h x = \<one>\<^bsub>H\<^esub>}"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1115
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1116
lemma (in group_hom) subgroup_kernel: "subgroup (kernel G H h) G"
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1117
  by (auto simp add: kernel_def group.intro intro: subgroup.intro)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1118
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1119
text\<open>The kernel of a homomorphism is a normal subgroup\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1120
lemma (in group_hom) normal_kernel: "(kernel G H h) \<lhd> G"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1121
  apply (simp only: G.normal_inv_iff subgroup_kernel)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1122
  apply (simp add: kernel_def)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1123
  done
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1124
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1125
lemma iso_kernel_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1126
  assumes "group G" "group H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1127
  shows "f \<in> iso G H \<longleftrightarrow> f \<in> hom G H \<and> kernel G H f = {\<one>\<^bsub>G\<^esub>} \<and> f ` carrier G = carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1128
    (is "?lhs = ?rhs")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1129
proof (intro iffI conjI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1130
  assume f: ?lhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1131
  show "f \<in> hom G H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1132
    using Group.iso_iff f by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1133
  show "kernel G H f = {\<one>\<^bsub>G\<^esub>}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1134
    using assms f Group.group_def hom_one
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1135
    by (fastforce simp add: kernel_def iso_iff_mon_epi mon_iff_hom_one set_eq_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1136
  show "f ` carrier G = carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1137
    by (meson Group.iso_iff f)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1138
next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1139
  assume ?rhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1140
  with assms show ?lhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1141
    by (auto simp: kernel_def iso_def bij_betw_def inj_on_one_iff')
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1142
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1143
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1144
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1145
lemma (in group_hom) FactGroup_nonempty:
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1146
  assumes "X \<in> carrier (G Mod kernel G H h)"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1147
  shows "X \<noteq> {}"
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1148
  using assms unfolding FactGroup_def
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1149
  by (metis group_hom.subgroup_kernel group_hom_axioms partial_object.simps(1) subgroup.rcosets_non_empty)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1150
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1151
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1152
lemma (in group_hom) FactGroup_universal_kernel:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1153
  assumes "N \<lhd> G" and h: "N \<subseteq> kernel G H h"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1154
  obtains g where "g \<in> hom (G Mod N) H" "\<And>x. x \<in> carrier G \<Longrightarrow> g(r_coset G N x) = h x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1155
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1156
  have "h x = h y"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1157
    if "x \<in> carrier G" "y \<in> carrier G" "r_coset G N x = r_coset G N y" for x y
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1158
  proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1159
    have "x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y \<in> N"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1160
      using \<open>N \<lhd> G\<close> group.rcos_self normal.axioms(2) normal_imp_subgroup
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1161
         subgroup.rcos_module_imp that by metis 
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1162
    with h have xy: "x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y \<in> kernel G H h"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1163
      by blast
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1164
    have "h x \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub>(h y) = h (x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1165
      by (simp add: that)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1166
    also have "\<dots> = \<one>\<^bsub>H\<^esub>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1167
      using xy by (simp add: kernel_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1168
    finally have "h x \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub>(h y) = \<one>\<^bsub>H\<^esub>" .
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1169
    then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1170
      using H.inv_equality that by fastforce
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1171
  qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1172
  with FactGroup_universal [OF homh \<open>N \<lhd> G\<close>] that show thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1173
    by metis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1174
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1175
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1176
lemma (in group_hom) FactGroup_the_elem_mem:
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1177
  assumes X: "X \<in> carrier (G Mod (kernel G H h))"
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1178
  shows "the_elem (h`X) \<in> carrier H"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1179
proof -
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1180
  from X
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1181
  obtain g where g: "g \<in> carrier G"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1182
             and "X = kernel G H h #> g"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1183
    by (auto simp add: FactGroup_def RCOSETS_def)
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61628
diff changeset
  1184
  hence "h ` X = {h g}" by (auto simp add: kernel_def r_coset_def g intro!: imageI)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1185
  thus ?thesis by (auto simp add: g)
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1186
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1187
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1188
lemma (in group_hom) FactGroup_hom:
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1189
     "(\<lambda>X. the_elem (h`X)) \<in> hom (G Mod (kernel G H h)) H"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1190
proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1191
  have "the_elem (h ` (X <#> X')) = the_elem (h ` X) \<otimes>\<^bsub>H\<^esub> the_elem (h ` X')"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1192
    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'
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1193
  proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1194
    obtain g and g'
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1195
      where "g \<in> carrier G" and "g' \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1196
        and "X = kernel G H h #> g" and "X' = kernel G H h #> g'"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1197
      using X X' by (auto simp add: FactGroup_def RCOSETS_def)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1198
    hence all: "\<forall>x\<in>X. h x = h g" "\<forall>x\<in>X'. h x = h g'"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1199
      and Xsub: "X \<subseteq> carrier G" and X'sub: "X' \<subseteq> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1200
      by (force simp add: kernel_def r_coset_def image_def)+
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1201
    hence "h ` (X <#> X') = {h g \<otimes>\<^bsub>H\<^esub> h g'}" using X X'
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1202
      by (auto dest!: FactGroup_nonempty intro!: image_eqI
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1203
          simp add: set_mult_def
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1204
          subsetD [OF Xsub] subsetD [OF X'sub])
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1205
    then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \<otimes>\<^bsub>H\<^esub> the_elem (h ` X')"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1206
      by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1207
  qed
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1208
  then show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1209
    by (simp add: hom_def FactGroup_the_elem_mem normal.factorgroup_is_group [OF normal_kernel] group.axioms monoid.m_closed)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1210
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1211
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1212
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1213
text\<open>Lemma for the following injectivity result\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1214
lemma (in group_hom) FactGroup_subset:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1215
  assumes "g \<in> carrier G" "g' \<in> carrier G" "h g = h g'"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1216
  shows "kernel G H h #> g \<subseteq> kernel G H h #> g'"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1217
  unfolding kernel_def r_coset_def
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1218
proof clarsimp
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1219
  fix y 
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1220
  assume "y \<in> carrier G" "h y = \<one>\<^bsub>H\<^esub>"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1221
  with assms show "\<exists>x. x \<in> carrier G \<and> h x = \<one>\<^bsub>H\<^esub> \<and> y \<otimes> g = x \<otimes> g'"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1222
    by (rule_tac x="y \<otimes> g \<otimes> inv g'" in exI) (auto simp: G.m_assoc)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1223
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1224
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1225
lemma (in group_hom) FactGroup_inj_on:
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1226
     "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: 68517
diff changeset
  1227
proof (simp add: inj_on_def, clarify)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1228
  fix X and X'
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1229
  assume X:  "X  \<in> carrier (G Mod kernel G H h)"
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1230
     and X': "X' \<in> carrier (G Mod kernel G H h)"
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1231
  then
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1232
  obtain g and g'
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1233
           where gX: "g \<in> carrier G"  "g' \<in> carrier G"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1234
              "X = kernel G H h #> g" "X' = kernel G H h #> g'"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1235
    by (auto simp add: FactGroup_def RCOSETS_def)
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1236
  hence all: "\<forall>x\<in>X. h x = h g" "\<forall>x\<in>X'. h x = h g'"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1237
    by (force simp add: kernel_def r_coset_def image_def)+
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1238
  assume "the_elem (h ` X) = the_elem (h ` X')"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1239
  hence h: "h g = h g'"
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1240
    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: 68517
diff changeset
  1241
  show "X=X'" by (rule equalityI) (simp_all add: FactGroup_subset h gX)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1242
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1243
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
  1244
text\<open>If the homomorphism \<^term>\<open>h\<close> is onto \<^term>\<open>H\<close>, then so is the
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1245
homomorphism from the quotient group\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1246
lemma (in group_hom) FactGroup_onto:
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1247
  assumes h: "h ` carrier G = carrier H"
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1248
  shows "(\<lambda>X. the_elem (h ` X)) ` carrier (G Mod kernel G H h) = carrier H"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1249
proof
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1250
  show "(\<lambda>X. the_elem (h ` X)) ` carrier (G Mod kernel G H h) \<subseteq> carrier H"
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1251
    by (auto simp add: FactGroup_the_elem_mem)
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1252
  show "carrier H \<subseteq> (\<lambda>X. the_elem (h ` X)) ` carrier (G Mod kernel G H h)"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1253
  proof
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1254
    fix y
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1255
    assume y: "y \<in> carrier H"
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1256
    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: 68517
diff changeset
  1257
      by (blast elim: equalityE)
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1258
    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: 68517
diff changeset
  1259
      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: 68517
diff changeset
  1260
    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: 61628
diff changeset
  1261
      apply (auto intro!: bexI image_eqI simp add: FactGroup_def RCOSETS_def)
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61628
diff changeset
  1262
      apply (subst the_elem_image_unique)
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61628
diff changeset
  1263
      apply auto
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61628
diff changeset
  1264
      done
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1265
  qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1266
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1267
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1268
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
  1269
text\<open>If \<^term>\<open>h\<close> is a homomorphism from \<^term>\<open>G\<close> onto \<^term>\<open>H\<close>, then the
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
  1270
 quotient group \<^term>\<open>G Mod (kernel G H h)\<close> is isomorphic to \<^term>\<open>H\<close>.\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1271
theorem (in group_hom) FactGroup_iso_set:
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1272
  "h ` carrier G = carrier H
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1273
   \<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: 68517
diff changeset
  1274
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: 68517
diff changeset
  1275
              FactGroup_onto)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1276
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1277
corollary (in group_hom) FactGroup_iso :
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1278
  "h ` carrier G = carrier H
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1279
   \<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: 67443
diff changeset
  1280
  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: 67443
diff changeset
  1281
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1282
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1283
lemma (in group_hom) trivial_hom_iff: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1284
  "h ` (carrier G) = { \<one>\<^bsub>H\<^esub> } \<longleftrightarrow> kernel G H h = carrier G"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1285
  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: 67443
diff changeset
  1286
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1287
lemma (in group_hom) trivial_ker_imp_inj: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1288
  assumes "kernel G H h = { \<one> }"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1289
  shows "inj_on h (carrier G)"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1290
proof (rule inj_onI)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1291
  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: 67443
diff changeset
  1292
  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: 67443
diff changeset
  1293
  hence "g1 \<otimes> (inv g2) = \<one>"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1294
    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: 67443
diff changeset
  1295
  thus "g1 = g2"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1296
    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: 67443
diff changeset
  1297
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1298
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1299
lemma (in group_hom) inj_iff_trivial_ker:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1300
  shows "inj_on h (carrier G) \<longleftrightarrow> kernel G H h = { \<one> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1301
proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1302
  assume inj: "inj_on h (carrier G)" show "kernel G H h = { \<one> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1303
    unfolding kernel_def
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1304
  proof (auto)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1305
    fix a assume "a \<in> carrier G" "h a = \<one>\<^bsub>H\<^esub>" thus "a = \<one>"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1306
      using inj hom_one unfolding inj_on_def by force
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1307
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1308
next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1309
  show "kernel G H h = { \<one> } \<Longrightarrow> inj_on h (carrier G)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1310
    using trivial_ker_imp_inj by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1311
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1312
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1313
lemma (in group_hom) induced_group_hom':
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1314
  assumes "subgroup I G" shows "group_hom (G \<lparr> carrier := I \<rparr>) H h"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1315
proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1316
  have "h \<in> hom (G \<lparr> carrier := I \<rparr>) H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1317
    using homh subgroup.subset[OF assms] unfolding hom_def by (auto, meson hom_mult subsetCE)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1318
  thus ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1319
    using subgroup.subgroup_is_group[OF assms G.group_axioms] group_axioms
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1320
    unfolding group_hom_def group_hom_axioms_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1321
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1322
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1323
lemma (in group_hom) inj_on_subgroup_iff_trivial_ker:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1324
  assumes "subgroup I G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1325
  shows "inj_on h I \<longleftrightarrow> kernel (G \<lparr> carrier := I \<rparr>) H h = { \<one> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1326
  using group_hom.inj_iff_trivial_ker[OF induced_group_hom'[OF assms]] by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1327
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1328
lemma set_mult_hom:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1329
  assumes "h \<in> hom G H" "I \<subseteq> carrier G" and "J \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1330
  shows "h ` (I <#>\<^bsub>G\<^esub> J) = (h ` I) <#>\<^bsub>H\<^esub> (h ` J)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1331
proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1332
  show "h ` (I <#>\<^bsub>G\<^esub> J) \<subseteq> (h ` I) <#>\<^bsub>H\<^esub> (h ` J)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1333
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1334
    fix a assume "a \<in> h ` (I <#>\<^bsub>G\<^esub> J)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1335
    then obtain i j where i: "i \<in> I" and j: "j \<in> J" and "a = h (i \<otimes>\<^bsub>G\<^esub> j)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1336
      unfolding set_mult_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1337
    hence "a = (h i) \<otimes>\<^bsub>H\<^esub> (h j)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1338
      using assms unfolding hom_def by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1339
    thus "a \<in> (h ` I) <#>\<^bsub>H\<^esub> (h ` J)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1340
      using i and j unfolding set_mult_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1341
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1342
next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1343
  show "(h ` I) <#>\<^bsub>H\<^esub> (h ` J) \<subseteq> h ` (I <#>\<^bsub>G\<^esub> J)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1344
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1345
    fix a assume "a \<in> (h ` I) <#>\<^bsub>H\<^esub> (h ` J)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1346
    then obtain i j where i: "i \<in> I" and j: "j \<in> J" and "a = (h i) \<otimes>\<^bsub>H\<^esub> (h j)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1347
      unfolding set_mult_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1348
    hence "a = h (i \<otimes>\<^bsub>G\<^esub> j)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1349
      using assms unfolding hom_def by fastforce
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1350
    thus "a \<in> h ` (I <#>\<^bsub>G\<^esub> J)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1351
      using i and j unfolding set_mult_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1352
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1353
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1354
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1355
corollary coset_hom:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1356
  assumes "h \<in> hom G H" "I \<subseteq> carrier G" "a \<in> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1357
  shows "h ` (a <#\<^bsub>G\<^esub> I) = h a <#\<^bsub>H\<^esub> (h ` I)" and "h ` (I #>\<^bsub>G\<^esub> a) = (h ` I) #>\<^bsub>H\<^esub> h a"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1358
  unfolding l_coset_eq_set_mult r_coset_eq_set_mult using assms set_mult_hom[OF assms(1)] by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1359
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1360
corollary (in group_hom) set_mult_ker_hom:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1361
  assumes "I \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1362
  shows "h ` (I <#> (kernel G H h)) = h ` I" and "h ` ((kernel G H h) <#> I) = h ` I"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1363
proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1364
  have ker_in_carrier: "kernel G H h \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1365
    unfolding kernel_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1366
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1367
  have "h ` (kernel G H h) = { \<one>\<^bsub>H\<^esub> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1368
    unfolding kernel_def by force
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1369
  moreover have "h ` I \<subseteq> carrier H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1370
    using assms by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1371
  hence "(h ` I) <#>\<^bsub>H\<^esub> { \<one>\<^bsub>H\<^esub> } = h ` I" and "{ \<one>\<^bsub>H\<^esub> } <#>\<^bsub>H\<^esub> (h ` I) = h ` I"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1372
    unfolding set_mult_def by force+
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1373
  ultimately show "h ` (I <#> (kernel G H h)) = h ` I" and "h ` ((kernel G H h) <#> I) = h ` I"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1374
    using set_mult_hom[OF homh assms ker_in_carrier] set_mult_hom[OF homh ker_in_carrier assms] by simp+
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1375
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1376
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1377
subsubsection\<open>Trivial homomorphisms\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1378
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1379
definition trivial_homomorphism where
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1380
 "trivial_homomorphism G H f \<equiv> f \<in> hom G H \<and> (\<forall>x \<in> carrier G. f x = one H)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1381
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1382
lemma trivial_homomorphism_kernel:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1383
   "trivial_homomorphism G H f \<longleftrightarrow> f \<in> hom G H \<and> kernel G H f = carrier G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1384
  by (auto simp: trivial_homomorphism_def kernel_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1385
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1386
lemma (in group) trivial_homomorphism_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1387
   "trivial_homomorphism G H f \<longleftrightarrow> f \<in> hom G H \<and> f ` carrier G = {one H}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1388
  by (auto simp: trivial_homomorphism_def) (metis one_closed rev_image_eqI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1389
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1390
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1391
subsection \<open>Image kernel theorems\<close>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1392
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1393
lemma group_Int_image_ker:
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1394
  assumes f: "f \<in> hom G H" and g: "g \<in> hom H K" 
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1395
    and "inj_on (g \<circ> f) (carrier G)" "group G" "group H" "group K"
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1396
  shows "(f ` carrier G) \<inter> (kernel H K g) = {\<one>\<^bsub>H\<^esub>}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1397
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1398
  have "(f ` carrier G) \<inter> (kernel H K g) \<subseteq> {\<one>\<^bsub>H\<^esub>}"
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1399
    using assms 
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1400
    apply (clarsimp simp: kernel_def o_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1401
    by (metis group.is_monoid hom_one inj_on_eq_iff monoid.one_closed)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1402
  moreover have "one H \<in> f ` carrier G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1403
    by (metis f \<open>group G\<close> \<open>group H\<close> group.is_monoid hom_one image_iff monoid.one_closed)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1404
  moreover have "one H \<in> kernel H K g"
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1405
    unfolding kernel_def using Group.group_def \<open>group H\<close> \<open>group K\<close> g hom_one by blast
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1406
  ultimately show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1407
    by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1408
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1409
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1410
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1411
lemma group_sum_image_ker:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1412
  assumes f: "f \<in> hom G H" and g: "g \<in> hom H K" and eq: "(g \<circ> f) ` (carrier G) = carrier K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1413
     and "group G" "group H" "group K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1414
  shows "set_mult H (f ` carrier G) (kernel H K g) = carrier H" (is "?lhs = ?rhs")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1415
proof
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1416
  show "?lhs \<subseteq> ?rhs"
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1417
    apply (clarsimp simp: kernel_def set_mult_def)
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1418
    by (meson \<open>group H\<close> f group.is_monoid hom_in_carrier monoid.m_closed)
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1419
  have "\<exists>x\<in>carrier G. \<exists>z. z \<in> carrier H \<and> g z = \<one>\<^bsub>K\<^esub> \<and> y = f x \<otimes>\<^bsub>H\<^esub> z"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1420
    if y: "y \<in> carrier H" for y
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1421
  proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1422
    have "g y \<in> carrier K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1423
      using g hom_carrier that by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1424
    with assms obtain x where x: "x \<in> carrier G" "(g \<circ> f) x = g y"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1425
      by (metis image_iff)
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1426
    with assms have invf: "inv\<^bsub>H\<^esub> f x \<otimes>\<^bsub>H\<^esub> y \<in> carrier H"
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1427
      by (metis group.subgroup_self hom_carrier image_subset_iff subgroup_def y)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1428
    moreover
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1429
    have "g (inv\<^bsub>H\<^esub> f x \<otimes>\<^bsub>H\<^esub> y) = \<one>\<^bsub>K\<^esub>"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1430
    proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1431
      have "inv\<^bsub>H\<^esub> f x \<in> carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1432
        by (meson \<open>group H\<close> f group.inv_closed hom_carrier image_subset_iff x(1))
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1433
      then have "g (inv\<^bsub>H\<^esub> f x \<otimes>\<^bsub>H\<^esub> y) = g (inv\<^bsub>H\<^esub> f x) \<otimes>\<^bsub>K\<^esub> g y"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1434
        by (simp add: hom_mult [OF g] y)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1435
      also have "\<dots> = inv\<^bsub>K\<^esub> (g (f x)) \<otimes>\<^bsub>K\<^esub> g y"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1436
        using assms x(1)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1437
        by (metis (mono_tags, lifting) group_hom.hom_inv group_hom.intro group_hom_axioms.intro hom_carrier image_subset_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1438
      also have "\<dots> = \<one>\<^bsub>K\<^esub>"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1439
        using \<open>g y \<in> carrier K\<close> assms(6) group.l_inv x(2) by fastforce
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1440
      finally show ?thesis .
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1441
    qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1442
    moreover
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1443
    have "y = f x \<otimes>\<^bsub>H\<^esub> (inv\<^bsub>H\<^esub> f x \<otimes>\<^bsub>H\<^esub> y)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1444
      using x y
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1445
      by (meson \<open>group H\<close> invf f group.inv_solve_left' hom_in_carrier)
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1446
    ultimately
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1447
    show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1448
      using x y by force
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1449
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1450
  then show "?rhs \<subseteq> ?lhs"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1451
    by (auto simp: kernel_def set_mult_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1452
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1453
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1454
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1455
lemma group_sum_ker_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1456
  assumes f: "f \<in> hom G H" and g: "g \<in> hom H K" and eq: "(g \<circ> f) ` (carrier G) = carrier K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1457
     and "group G" "group H" "group K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1458
   shows "set_mult H (kernel H K g) (f ` carrier G) = carrier H" (is "?lhs = ?rhs")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1459
proof
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1460
  show "?lhs \<subseteq> ?rhs"
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1461
    apply (clarsimp simp: kernel_def set_mult_def)
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1462
    by (meson \<open>group H\<close> f group.is_monoid hom_in_carrier monoid.m_closed)
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1463
  have "\<exists>w\<in>carrier H. \<exists>x \<in> carrier G. g w = \<one>\<^bsub>K\<^esub> \<and> y = w \<otimes>\<^bsub>H\<^esub> f x"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1464
    if y: "y \<in> carrier H" for y
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1465
  proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1466
    have "g y \<in> carrier K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1467
      using g hom_carrier that by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1468
    with assms obtain x where x: "x \<in> carrier G" "(g \<circ> f) x = g y"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1469
      by (metis image_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1470
    with assms have carr: "(y \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> f x) \<in> carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1471
      by (metis group.subgroup_self hom_carrier image_subset_iff subgroup_def y)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1472
    moreover
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1473
    have "g (y \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> f x) = \<one>\<^bsub>K\<^esub>"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1474
    proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1475
      have "inv\<^bsub>H\<^esub> f x \<in> carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1476
        by (meson \<open>group H\<close> f group.inv_closed hom_carrier image_subset_iff x(1))
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1477
      then have "g (y \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> f x) = g y \<otimes>\<^bsub>K\<^esub> g (inv\<^bsub>H\<^esub> f x)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1478
        by (simp add: hom_mult [OF g] y)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1479
      also have "\<dots> = g y \<otimes>\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> (g (f x))"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1480
        using assms x(1)
77362
1a6103f6ab0b tidying ugly proofs
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1481
        by (metis group_hom.hom_inv group_hom_axioms.intro group_hom_def hom_in_carrier)
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1482
      also have "\<dots> = \<one>\<^bsub>K\<^esub>"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1483
        using \<open>g y \<in> carrier K\<close> assms(6) group.l_inv x(2)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1484
        by (simp add: group.r_inv)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1485
      finally show ?thesis .
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1486
    qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1487
    moreover
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1488
    have "y = (y \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> f x) \<otimes>\<^bsub>H\<^esub> f x"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1489
      using x y by (meson \<open>group H\<close> carr f group.inv_solve_right hom_carrier image_subset_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1490
    ultimately
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1491
    show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1492
      using x y by force
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1493
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1494
  then show "?rhs \<subseteq> ?lhs"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1495
    by (force simp: kernel_def set_mult_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1496
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1497
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1498
lemma group_semidirect_sum_ker_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1499
  assumes "(g \<circ> f) \<in> iso G K" "f \<in> hom G H" "g \<in> hom H K" "group G" "group H" "group K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1500
  shows "(kernel H K g) \<inter> (f ` carrier G) = {\<one>\<^bsub>H\<^esub>}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1501
        "kernel H K g <#>\<^bsub>H\<^esub> (f ` carrier G) = carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1502
  using assms
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1503
  by (simp_all add: iso_iff_mon_epi group_Int_image_ker group_sum_ker_image epi_def mon_def Int_commute [of "kernel H K g"])
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1504
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1505
lemma group_semidirect_sum_image_ker:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1506
  assumes f: "f \<in> hom G H" and g: "g \<in> hom H K" and iso: "(g \<circ> f) \<in> iso G K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1507
     and "group G" "group H" "group K"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1508
   shows "(f ` carrier G) \<inter> (kernel H K g) = {\<one>\<^bsub>H\<^esub>}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1509
          "f ` carrier G <#>\<^bsub>H\<^esub> (kernel H K g) = carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1510
  using group_Int_image_ker [OF f g] group_sum_image_ker [OF f g] assms
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1511
  by (simp_all add: iso_def bij_betw_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1512
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1513
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1514
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1515
subsection \<open>Factor Groups and Direct product\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1516
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1517
lemma (in group) DirProd_normal : \<^marker>\<open>contributor \<open>Martin Baillon\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1518
  assumes "group K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1519
    and "H \<lhd> G"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1520
    and "N \<lhd> K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1521
  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: 67443
diff changeset
  1522
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: 67443
diff changeset
  1523
  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: 67443
diff changeset
  1524
    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: 67443
diff changeset
  1525
         normal_imp_subgroup[OF assms(3)]].
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1526
  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: 67443
diff changeset
  1527
  proof-
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1528
    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: 68445
diff changeset
  1529
    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: 67443
diff changeset
  1530
    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: 67443
diff changeset
  1531
      unfolding DirProd_def by fastforce
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1532
    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: 67443
diff changeset
  1533
      unfolding DirProd_def by fastforce
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1534
    hence h1h2GK : "h1 \<in> carrier G" "h2 \<in> carrier K"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1535
      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: 67443
diff changeset
  1536
    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: 67443
diff changeset
  1537
      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: 67443
diff changeset
  1538
    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: 67443
diff changeset
  1539
      using h1h2 x1x2 h1h2GK by auto
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1540
    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
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1541
      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: 67443
diff changeset
  1542
    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: 67443
diff changeset
  1543
    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: 67443
diff changeset
  1544
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1545
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1546
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1547
lemma (in group) FactGroup_DirProd_multiplication_iso_set : \<^marker>\<open>contributor \<open>Martin Baillon\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1548
  assumes "group K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1549
    and "H \<lhd> G"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1550
    and "N \<lhd> K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1551
  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: 67443
diff changeset
  1552
proof-
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1553
  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
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1554
    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: 67443
diff changeset
  1555
  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: 67443
diff changeset
  1556
                \<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
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
  1557
    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: 67443
diff changeset
  1558
  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: 67443
diff changeset
  1559
                 \<forall>ya\<in>carrier (K Mod N).  x \<times> y = xa \<times> ya \<longrightarrow> x = xa \<and> y = ya)"
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68452
diff changeset
  1560
    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: 67443
diff changeset
  1561
    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: 68517
diff changeset
  1562
  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: 67443
diff changeset
  1563
                                     carrier (G \<times>\<times> K Mod H \<times> N)"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1564
  proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1565
    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)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1566
      using R by force
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1567
    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"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1568
      unfolding DirProd_def FactGroup_def RCOSETS_def r_coset_def by force
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1569
    show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1570
      unfolding image_def by (auto simp: intro: 1 2)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1571
  qed
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1572
  ultimately show ?thesis
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1573
    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: 67443
diff changeset
  1574
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1575
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1576
corollary (in group) FactGroup_DirProd_multiplication_iso_1 : \<^marker>\<open>contributor \<open>Martin Baillon\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1577
  assumes "group K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1578
    and "H \<lhd> G"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1579
    and "N \<lhd> K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1580
  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: 67443
diff changeset
  1581
  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: 67443
diff changeset
  1582
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1583
corollary (in group) FactGroup_DirProd_multiplication_iso_2 : \<^marker>\<open>contributor \<open>Martin Baillon\<close>\<close>
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1584
  assumes "group K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1585
    and "H \<lhd> G"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1586
    and "N \<lhd> K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1587
  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: 67443
diff changeset
  1588
  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: 67443
diff changeset
  1589
        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: 67443
diff changeset
  1590
  by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1591
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: 68443
diff changeset
  1592
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: 68443
diff changeset
  1593
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: 77362
diff changeset
  1594
text \<open>A group multiplied by a subgroup stays the same\<close>
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: 68443
diff changeset
  1595
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: 68443
diff changeset
  1596
  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: 68443
diff changeset
  1597
  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: 68443
diff changeset
  1598
proof
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1599
  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: 68445
diff changeset
  1600
    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: 68443
diff changeset
  1601
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: 68443
diff changeset
  1602
  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: 68443
diff changeset
  1603
  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: 68443
diff changeset
  1604
    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: 68443
diff changeset
  1605
  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: 68443
diff changeset
  1606
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: 68443
diff changeset
  1607
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: 77362
diff changeset
  1608
text \<open>Same lemma as above, but everything is included in a subgroup\<close>
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: 68443
diff changeset
  1609
lemma (in group) set_mult_subgroup_idem:
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1610
  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: 68517
diff changeset
  1611
  shows "H <#> N = H"
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1612
  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: 68443
diff changeset
  1613
77407
02af8a1b97f6 Fixed a presentation error
paulson <lp15@cam.ac.uk>
parents: 77406
diff changeset
  1614
text \<open>A normal subgroup is commutative with set multiplication\<close>
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: 68443
diff changeset
  1615
lemma (in group) commut_normal:
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1616
  assumes "subgroup H G" and "N\<lhd>G"
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1617
  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: 68443
diff changeset
  1618
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: 68443
diff changeset
  1619
  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: 68443
diff changeset
  1620
  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: 68443
diff changeset
  1621
  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: 68443
diff changeset
  1622
  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: 68443
diff changeset
  1623
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: 68443
diff changeset
  1624
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: 77362
diff changeset
  1625
text \<open>Same lemma as above, but everything is included in a subgroup\<close>
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: 68443
diff changeset
  1626
lemma (in group) commut_normal_subgroup:
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1627
  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: 68517
diff changeset
  1628
    and "subgroup K (G \<lparr> carrier := H \<rparr>)"
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1629
  shows "K <#> N = N <#> K"
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: 77362
diff changeset
  1630
  by (metis assms(2) assms(3) group.commut_normal normal.axioms(2) set_mult_consistent)
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: 68443
diff changeset
  1631
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1632
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1633
subsubsection "Lemmas about intersection and normal subgroups"
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: 77362
diff changeset
  1634
text \<open>Mostly by Jakob von Raumer\<close>
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: 68443
diff changeset
  1635
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1636
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: 68443
diff changeset
  1637
  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: 68443
diff changeset
  1638
    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: 68443
diff changeset
  1639
    and "H1\<lhd>G\<lparr>carrier := H\<rparr>"
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: 77362
diff changeset
  1640
  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: 68443
diff changeset
  1641
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: 68443
diff changeset
  1642
  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: 68443
diff changeset
  1643
    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: 68443
diff changeset
  1644
  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: 68443
diff changeset
  1645
  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: 68443
diff changeset
  1646
    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: 68443
diff changeset
  1647
      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: 68443
diff changeset
  1648
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1649
  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: 68443
diff changeset
  1650
    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: 68443
diff changeset
  1651
    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: 68443
diff changeset
  1652
      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: 68443
diff changeset
  1653
        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: 68443
diff changeset
  1654
    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: 68443
diff changeset
  1655
      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: 68443
diff changeset
  1656
    next
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1657
      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: 68445
diff changeset
  1658
        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: 68443
diff changeset
  1659
      also have "... \<subseteq> H" by simp
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1660
      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: 68443
diff changeset
  1661
        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: 68443
diff changeset
  1662
    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: 68443
diff changeset
  1663
    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: 68443
diff changeset
  1664
  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: 68443
diff changeset
  1665
    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: 68443
diff changeset
  1666
    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: 68443
diff changeset
  1667
      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: 68517
diff changeset
  1668
        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: 68443
diff changeset
  1669
      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: 68443
diff changeset
  1670
        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: 68443
diff changeset
  1671
      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: 68443
diff changeset
  1672
      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: 68443
diff changeset
  1673
      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: 68443
diff changeset
  1674
        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: 68443
diff changeset
  1675
      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: 68443
diff changeset
  1676
      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: 68517
diff changeset
  1677
        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: 68443
diff changeset
  1678
      have "H1\<subseteq>carrier(GH)"
68452
c027dfbfad30 more on infinite products. Also subgroup_imp_subset -> subgroup.subset
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
  1679
        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: 68517
diff changeset
  1680
      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: 68443
diff changeset
  1681
        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: 68443
diff changeset
  1682
      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: 68443
diff changeset
  1683
        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: 68517
diff changeset
  1684
      have xH:"x\<in>carrier(GH)"
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1685
        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: 68443
diff changeset
  1686
      have hH:"h\<in>carrier(GH)"
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1687
        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: 68443
diff changeset
  1688
      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
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1689
        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: 68443
diff changeset
  1690
      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: 68443
diff changeset
  1691
        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: 68443
diff changeset
  1692
      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: 68443
diff changeset
  1693
        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: 68443
diff changeset
  1694
        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: 68443
diff changeset
  1695
      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: 68443
diff changeset
  1696
      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: 68443
diff changeset
  1697
      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: 68443
diff changeset
  1698
    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: 68443
diff changeset
  1699
  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: 68443
diff changeset
  1700
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: 68443
diff changeset
  1701
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69895
diff changeset
  1702
lemma (in group) normal_Int_subgroup:
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: 68443
diff changeset
  1703
  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: 68443
diff changeset
  1704
    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: 68443
diff changeset
  1705
  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: 68443
diff changeset
  1706
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: 68443
diff changeset
  1707
  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: 68443
diff changeset
  1708
  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: 68443
diff changeset
  1709
  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: 68443
diff changeset
  1710
  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: 68443
diff changeset
  1711
  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: 68443
diff changeset
  1712
    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: 68445
diff changeset
  1713
  moreover have "K \<inter> H = H" using K_def assms subgroup.subset by blast
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1714
  ultimately show "normal (N\<inter>H) (G\<lparr>carrier := H\<rparr>)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1715
 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: 68443
diff changeset
  1716
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: 68443
diff changeset
  1717
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: 77362
diff changeset
  1718
lemma (in group) normal_restrict_supergroup:
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: 77362
diff changeset
  1719
  assumes "subgroup S G" "N \<lhd> G" "N \<subseteq> 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: 77362
diff changeset
  1720
  shows "N \<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: 77362
diff changeset
  1721
  by (metis assms inf.absorb_iff1 normal_Int_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: 77362
diff changeset
  1722
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: 77362
diff changeset
  1723
text \<open>A subgroup relation survives factoring by a normal subgroup.\<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: 77362
diff changeset
  1724
lemma (in group) normal_subgroup_factorize:
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: 77362
diff changeset
  1725
  assumes "N \<lhd> G" and "N \<subseteq> H" and "subgroup 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: 77362
diff changeset
  1726
  shows "subgroup (rcosets\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> N) (G Mod 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: 77362
diff changeset
  1727
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: 77362
diff changeset
  1728
  interpret GModN: group "G Mod 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: 77362
diff changeset
  1729
    using assms(1) by (rule normal.factorgroup_is_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: 77362
diff changeset
  1730
  have "N \<lhd> G\<lparr>carrier := H\<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: 77362
diff changeset
  1731
    using assms by (metis normal_restrict_supergroup)
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: 77362
diff changeset
  1732
  hence grpHN: "group (G\<lparr>carrier := H\<rparr> Mod 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: 77362
diff changeset
  1733
    by (rule normal.factorgroup_is_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: 77362
diff changeset
  1734
  have "(<#>\<^bsub>G\<lparr>carrier:=H\<rparr>\<^esub>) = (\<lambda>U K. (\<Union>h\<in>U. \<Union>k\<in>K. {h \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> k}))" 
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: 77362
diff changeset
  1735
    using set_mult_def 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: 77362
diff changeset
  1736
  moreover have "\<dots> = (\<lambda>U K. (\<Union>h\<in>U. \<Union>k\<in>K. {h \<otimes>\<^bsub>G\<^esub> k}))" 
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: 77362
diff changeset
  1737
    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: 77362
diff changeset
  1738
  moreover have "(<#>) = (\<lambda>U K. (\<Union>h\<in>U. \<Union>k\<in>K. {h \<otimes> k}))" 
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: 77362
diff changeset
  1739
    using set_mult_def 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: 77362
diff changeset
  1740
  ultimately have "(<#>\<^bsub>G\<lparr>carrier:=H\<rparr>\<^esub>) = (<#>\<^bsub>G\<^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: 77362
diff changeset
  1741
    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: 77362
diff changeset
  1742
  with grpHN have "group ((G Mod N)\<lparr>carrier := (rcosets\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> N)\<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: 77362
diff changeset
  1743
    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: 77362
diff changeset
  1744
  moreover have "rcosets\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> N \<subseteq> carrier (G Mod 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: 77362
diff changeset
  1745
    unfolding FactGroup_def RCOSETS_def r_coset_def using assms(3) subgroup.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: 77362
diff changeset
  1746
    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: 77362
diff changeset
  1747
  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: 77362
diff changeset
  1748
    using GModN.group_incl_imp_subgroup 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: 77362
diff changeset
  1749
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: 77362
diff changeset
  1750
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: 77362
diff changeset
  1751
text \<open>A normality relation survives factoring by a normal subgroup.\<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: 77362
diff changeset
  1752
lemma (in group) normality_factorization:
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: 77362
diff changeset
  1753
  assumes NG: "N \<lhd> G" and NH: "N \<subseteq> H" and HG: "H \<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: 77362
diff changeset
  1754
  shows "(rcosets\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> N) \<lhd> (G Mod 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: 77362
diff changeset
  1755
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: 77362
diff changeset
  1756
  from assms(1) interpret GModN: group "G Mod N" by (metis normal.factorgroup_is_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: 77362
diff changeset
  1757
  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: 77362
diff changeset
  1758
    unfolding GModN.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: 77362
diff changeset
  1759
  proof (intro conjI strip)
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: 77362
diff changeset
  1760
    show "subgroup (rcosets\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> N) (G Mod 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: 77362
diff changeset
  1761
      using assms normal_imp_subgroup normal_subgroup_factorize 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: 77362
diff changeset
  1762
  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: 77362
diff changeset
  1763
    fix U V
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: 77362
diff changeset
  1764
    assume U: "U \<in> carrier (G Mod N)" and V: "V \<in> rcosets\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> 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: 77362
diff changeset
  1765
    then obtain g where g: "g \<in> carrier G" "U = 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: 77362
diff changeset
  1766
      unfolding FactGroup_def RCOSETS_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: 77362
diff changeset
  1767
    from V obtain h where h: "h \<in> H" "V = N #> 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: 77362
diff changeset
  1768
      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: 77362
diff changeset
  1769
    hence hG: "h \<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: 77362
diff changeset
  1770
      using HG normal_imp_subgroup 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: 77362
diff changeset
  1771
    hence ghG: "g \<otimes> h \<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: 77362
diff changeset
  1772
      using g m_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: 77362
diff changeset
  1773
    from g h have "g \<otimes> h \<otimes> inv 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: 77362
diff changeset
  1774
      using HG normal_inv_iff 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: 77362
diff changeset
  1775
    moreover have "U <#> V <#> inv\<^bsub>G Mod N\<^esub> U = N #> (g \<otimes> h \<otimes> inv 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: 77362
diff changeset
  1776
    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: 77362
diff changeset
  1777
      from g U have "inv\<^bsub>G Mod N\<^esub> U = N #> inv 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: 77362
diff changeset
  1778
        using NG normal.inv_FactGroup normal.rcos_inv 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: 77362
diff changeset
  1779
      hence "U <#> V <#> inv\<^bsub>G Mod N\<^esub> U = (N #> g) <#> (N #> h) <#> (N #> inv 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: 77362
diff changeset
  1780
        using g h 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: 77362
diff changeset
  1781
      also have "\<dots> = N #> (g \<otimes> h \<otimes> inv 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: 77362
diff changeset
  1782
        using g hG NG inv_closed ghG normal.rcos_sum 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: 77362
diff changeset
  1783
      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: 77362
diff changeset
  1784
    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: 77362
diff changeset
  1785
    ultimately show "U \<otimes>\<^bsub>G Mod N\<^esub> V \<otimes>\<^bsub>G Mod N\<^esub> inv\<^bsub>G Mod N\<^esub> U \<in> rcosets\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> 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: 77362
diff changeset
  1786
      unfolding 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: 77362
diff changeset
  1787
  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: 77362
diff changeset
  1788
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: 77362
diff changeset
  1789
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: 77362
diff changeset
  1790
text \<open>Factorizing by the trivial subgroup is an isomorphism.\<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: 77362
diff changeset
  1791
lemma (in group) trivial_factor_iso:
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: 77362
diff changeset
  1792
  shows "the_elem \<in> iso (G Mod {\<one>}) 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: 77362
diff changeset
  1793
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: 77362
diff changeset
  1794
  have "group_hom G G (\<lambda>x. 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: 77362
diff changeset
  1795
    unfolding group_hom_def group_hom_axioms_def hom_def using is_group 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: 77362
diff changeset
  1796
  moreover have "(\<lambda>x. x) ` carrier G = 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: 77362
diff changeset
  1797
    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: 77362
diff changeset
  1798
  moreover have "kernel G G (\<lambda>x. x) = {\<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: 77362
diff changeset
  1799
    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: 77362
diff changeset
  1800
  ultimately show ?thesis using 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: 77362
diff changeset
  1801
    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: 77362
diff changeset
  1802
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: 77362
diff changeset
  1803
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: 77362
diff changeset
  1804
text \<open>And the dual theorem to the previous one: Factorizing by the group itself gives the trivial group\<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: 77362
diff changeset
  1805
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: 77362
diff changeset
  1806
lemma (in group) self_factor_iso:
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: 77362
diff changeset
  1807
  shows "(\<lambda>X. the_elem ((\<lambda>x. \<one>) ` X)) \<in> iso (G Mod (carrier G)) (G\<lparr> carrier := {\<one>} \<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: 77362
diff changeset
  1808
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: 77362
diff changeset
  1809
  have "group (G\<lparr>carrier := {\<one>}\<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: 77362
diff changeset
  1810
    by (metis subgroup_imp_group triv_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: 77362
diff changeset
  1811
  hence "group_hom G (G\<lparr>carrier := {\<one>}\<rparr>) (\<lambda>x. \<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: 77362
diff changeset
  1812
    unfolding group_hom_def group_hom_axioms_def hom_def using is_group 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: 77362
diff changeset
  1813
  moreover have "(\<lambda>x. \<one>) ` carrier G = carrier (G\<lparr>carrier := {\<one>}\<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: 77362
diff changeset
  1814
    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: 77362
diff changeset
  1815
  moreover have "kernel G (G\<lparr>carrier := {\<one>}\<rparr>) (\<lambda>x. \<one>) = 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: 77362
diff changeset
  1816
    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: 77362
diff changeset
  1817
  ultimately show ?thesis using 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: 77362
diff changeset
  1818
    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: 77362
diff changeset
  1819
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: 77362
diff changeset
  1820
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: 77362
diff changeset
  1821
text \<open>Factoring by a normal subgroups yields the trivial group iff the subgroup is the whole group.\<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: 77362
diff changeset
  1822
lemma (in normal) fact_group_trivial_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: 77362
diff changeset
  1823
  assumes "finite (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: 77362
diff changeset
  1824
  shows "(carrier (G Mod H) = {\<one>\<^bsub>G Mod H\<^esub>}) \<longleftrightarrow> (H = 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: 77362
diff changeset
  1825
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: 77362
diff changeset
  1826
  assume "carrier (G Mod H) = {\<one>\<^bsub>G 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: 77362
diff changeset
  1827
  moreover have "order (G Mod H) * card H = order 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: 77362
diff changeset
  1828
    by (simp add: FactGroup_def lagrange order_def subgroup_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: 77362
diff changeset
  1829
  ultimately have "card H = order G" unfolding order_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: 77362
diff changeset
  1830
  thus "H = 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: 77362
diff changeset
  1831
    by (simp add: assms card_subset_eq order_def 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: 77362
diff changeset
  1832
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: 77362
diff changeset
  1833
  assume "H = 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: 77362
diff changeset
  1834
  with assms is_subgroup lagrange 
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: 77362
diff changeset
  1835
  have "card (rcosets H) * order G = order 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: 77362
diff changeset
  1836
    by (simp add: order_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: 77362
diff changeset
  1837
  then have "card (rcosets H) = 1" 
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: 77362
diff changeset
  1838
    using assms order_gt_0_iff_finite 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: 77362
diff changeset
  1839
  hence "order (G Mod H) = 1" 
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: 77362
diff changeset
  1840
    unfolding order_def 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: 77362
diff changeset
  1841
  thus "carrier (G Mod H) = {\<one>\<^bsub>G 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: 77362
diff changeset
  1842
    using factorgroup_is_group by (metis group.order_one_triv_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: 77362
diff changeset
  1843
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: 77362
diff changeset
  1844
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: 77362
diff changeset
  1845
text \<open>The union of all the cosets contained in a subgroup of a quotient group acts as a represenation for that subgroup.\<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: 77362
diff changeset
  1846
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: 77362
diff changeset
  1847
lemma (in normal) factgroup_subgroup_union_char:
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: 77362
diff changeset
  1848
  assumes "subgroup A (G 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: 77362
diff changeset
  1849
  shows "(\<Union>A) = {x \<in> carrier G. H #> x \<in> A}"
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: 77362
diff changeset
  1850
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: 77362
diff changeset
  1851
  show "\<Union>A \<subseteq> {x \<in> carrier G. H #> x \<in> A}"
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: 77362
diff changeset
  1852
  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: 77362
diff changeset
  1853
    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: 77362
diff changeset
  1854
    assume x: "x \<in> \<Union>A"
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: 77362
diff changeset
  1855
    then obtain a where a: "a \<in> A" "x \<in> a" and xx: "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: 77362
diff changeset
  1856
      using subgroup.subset assms by (force simp add: FactGroup_def RCOSETS_def r_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: 77362
diff changeset
  1857
    from assms a obtain y where y: "y \<in> carrier G" "a = 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: 77362
diff changeset
  1858
      using subgroup.subset unfolding FactGroup_def RCOSETS_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: 77362
diff changeset
  1859
    with a have "x \<in> H #> y" 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: 77362
diff changeset
  1860
    hence "H #> y = H #> x" using y is_subgroup repr_independence 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: 77362
diff changeset
  1861
    with y(2) a(1) have "H #> x \<in> A" 
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: 77362
diff changeset
  1862
      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: 77362
diff changeset
  1863
    with xx show "x \<in> {x \<in> carrier G. H #> x \<in> A}" 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: 77362
diff changeset
  1864
  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: 77362
diff changeset
  1865
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: 77362
diff changeset
  1866
  show "{x \<in> carrier G. H #> x \<in> A} \<subseteq> \<Union>A"
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: 77362
diff changeset
  1867
    using rcos_self subgroup_axioms 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: 77362
diff changeset
  1868
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: 77362
diff changeset
  1869
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: 77362
diff changeset
  1870
lemma (in normal) factgroup_subgroup_union_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: 77362
diff changeset
  1871
  assumes "subgroup A (G 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: 77362
diff changeset
  1872
  shows "subgroup (\<Union>A) 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: 77362
diff changeset
  1873
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: 77362
diff changeset
  1874
  have "subgroup {x \<in> carrier G. H #> x \<in> A} 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: 77362
diff changeset
  1875
  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: 77362
diff changeset
  1876
    show "{x \<in> carrier G. H #> x \<in> A} \<subseteq> carrier 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: 77362
diff changeset
  1877
  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: 77362
diff changeset
  1878
    fix x 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: 77362
diff changeset
  1879
    assume xy: "x \<in> {x \<in> carrier G. H #> x \<in> A}" "y \<in> {x \<in> carrier G. H #> x \<in> A}"
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: 77362
diff changeset
  1880
    then have "(H #> x) <#> (H #> y) \<in> A" 
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: 77362
diff changeset
  1881
      using subgroup.m_closed assms unfolding FactGroup_def  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: 77362
diff changeset
  1882
    hence "H #> (x \<otimes> y) \<in> A"
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: 77362
diff changeset
  1883
      using xy rcos_sum 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: 77362
diff changeset
  1884
    with xy show "x \<otimes> y \<in> {x \<in> carrier G. H #> x \<in> A}" 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: 77362
diff changeset
  1885
  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: 77362
diff changeset
  1886
    have "H #> \<one> \<in> A"
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: 77362
diff changeset
  1887
      using assms subgroup.one_closed subset 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: 77362
diff changeset
  1888
    with assms one_closed show "\<one> \<in> {x \<in> carrier G. H #> x \<in> A}" 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: 77362
diff changeset
  1889
  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: 77362
diff changeset
  1890
    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: 77362
diff changeset
  1891
    assume x: "x \<in> {x \<in> carrier G. H #> x \<in> A}"
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: 77362
diff changeset
  1892
    hence invx: "inv x \<in> carrier G" using inv_closed 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: 77362
diff changeset
  1893
    from assms x have "set_inv (H #> x) \<in> A" using subgroup.m_inv_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: 77362
diff changeset
  1894
      using inv_FactGroup subgroup.mem_carrier 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: 77362
diff changeset
  1895
    with invx show "inv x \<in> {x \<in> carrier G. H #> x \<in> A}"
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: 77362
diff changeset
  1896
      using rcos_inv 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: 77362
diff changeset
  1897
  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: 77362
diff changeset
  1898
  with assms factgroup_subgroup_union_char show ?thesis 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: 77362
diff changeset
  1899
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: 77362
diff changeset
  1900
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: 77362
diff changeset
  1901
lemma (in normal) factgroup_subgroup_union_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: 77362
diff changeset
  1902
  assumes "A \<lhd> (G 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: 77362
diff changeset
  1903
  shows "\<Union>A \<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: 77362
diff changeset
  1904
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: 77362
diff changeset
  1905
  have "{x \<in> carrier G. H #> x \<in> A} \<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: 77362
diff changeset
  1906
    unfolding normal_def normal_axioms_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: 77362
diff changeset
  1907
  proof (intro conjI strip)
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: 77362
diff changeset
  1908
    from assms show "subgroup {x \<in> carrier G. H #> x \<in> A} 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: 77362
diff changeset
  1909
      by (metis (full_types) factgroup_subgroup_union_char factgroup_subgroup_union_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: 77362
diff changeset
  1910
  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: 77362
diff changeset
  1911
    interpret Anormal: normal A "(G Mod H)" using assms 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: 77362
diff changeset
  1912
    show "{x \<in> carrier G. H #> x \<in> A} #> x = x <# {x \<in> carrier G. H #> x \<in> A}" if x: "x \<in> carrier G" for 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: 77362
diff changeset
  1913
    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: 77362
diff changeset
  1914
      { fix 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: 77362
diff changeset
  1915
        assume y: "y \<in> {x \<in> carrier G. H #> x \<in> A} #> 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: 77362
diff changeset
  1916
        then obtain x' where x': "x' \<in> carrier G" "H #> x' \<in> A" "y = x' \<otimes> 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: 77362
diff changeset
  1917
          unfolding 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: 77362
diff changeset
  1918
        from x(1) have Hx: "H #> x \<in> carrier (G 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: 77362
diff changeset
  1919
          unfolding FactGroup_def RCOSETS_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: 77362
diff changeset
  1920
        with x' have "(inv\<^bsub>G Mod H\<^esub> (H #> x)) \<otimes>\<^bsub>G Mod H\<^esub> (H #> x') \<otimes>\<^bsub>G Mod H\<^esub> (H #> x) \<in> A" 
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: 77362
diff changeset
  1921
          using Anormal.inv_op_closed1 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: 77362
diff changeset
  1922
        hence "(set_inv (H #> x)) <#> (H #> x') <#> (H #> x) \<in> A" 
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: 77362
diff changeset
  1923
          using inv_FactGroup Hx 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: 77362
diff changeset
  1924
        hence "(H #> (inv x)) <#> (H #> x') <#> (H #> x) \<in> A" 
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: 77362
diff changeset
  1925
          using x(1) by (metis rcos_inv)
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: 77362
diff changeset
  1926
        hence "H #> (inv x \<otimes> x' \<otimes> x) \<in> A" 
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: 77362
diff changeset
  1927
          by (metis inv_closed m_closed rcos_sum x'(1) x(1))
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: 77362
diff changeset
  1928
        moreover have "inv x \<otimes> x' \<otimes> 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: 77362
diff changeset
  1929
          using x x' by (metis inv_closed m_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: 77362
diff changeset
  1930
        ultimately have xcoset: "x \<otimes> (inv x \<otimes> x' \<otimes> x) \<in> x <# {x \<in> carrier G. H #> x \<in> A}" 
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: 77362
diff changeset
  1931
          unfolding l_coset_def using x(1) 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: 77362
diff changeset
  1932
        have "x \<otimes> (inv x \<otimes> x' \<otimes> x) = (x \<otimes> inv x) \<otimes> x' \<otimes> 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: 77362
diff changeset
  1933
          by (metis Units_eq Units_inv_Units m_assoc m_closed x'(1) x(1))
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: 77362
diff changeset
  1934
        also have "\<dots> = 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: 77362
diff changeset
  1935
          by (simp add: x 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: 77362
diff changeset
  1936
        finally have "x \<otimes> (inv x \<otimes> x' \<otimes> x) = 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: 77362
diff changeset
  1937
        with xcoset have "y \<in> x <# {x \<in> carrier G. H #> x \<in> A}" 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: 77362
diff changeset
  1938
      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: 77362
diff changeset
  1939
      { fix 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: 77362
diff changeset
  1940
        assume y: "y \<in> x <# {x \<in> carrier G. H #> x \<in> A}"
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: 77362
diff changeset
  1941
        then obtain x' where x': "x' \<in> carrier G" "H #> x' \<in> A" "y = x \<otimes> x'" unfolding 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: 77362
diff changeset
  1942
        from x(1) have invx: "inv 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: 77362
diff changeset
  1943
          by (rule inv_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: 77362
diff changeset
  1944
        hence Hinvx: "H #> (inv x) \<in> carrier (G 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: 77362
diff changeset
  1945
          unfolding FactGroup_def RCOSETS_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: 77362
diff changeset
  1946
        with x' have "(inv\<^bsub>G Mod H\<^esub> (H #> inv x)) \<otimes>\<^bsub>G Mod H\<^esub> (H #> x') \<otimes>\<^bsub>G Mod H\<^esub> (H #> inv x) \<in> A" 
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: 77362
diff changeset
  1947
          using invx Anormal.inv_op_closed1 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: 77362
diff changeset
  1948
        hence "(set_inv (H #> inv x)) <#> (H #> x') <#> (H #> inv x) \<in> A" 
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: 77362
diff changeset
  1949
          using inv_FactGroup Hinvx 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: 77362
diff changeset
  1950
        hence "H #> (x \<otimes> x' \<otimes> inv x) \<in> A"
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: 77362
diff changeset
  1951
          by (simp add: rcos_inv rcos_sum x x'(1))
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: 77362
diff changeset
  1952
        moreover have "x \<otimes> x' \<otimes> inv x \<in> carrier G" using x x' by (metis inv_closed m_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: 77362
diff changeset
  1953
        ultimately have xcoset: "(x \<otimes> x' \<otimes> inv x) \<otimes> x \<in> {x \<in> carrier G. H #> x \<in> A} #> 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: 77362
diff changeset
  1954
          unfolding r_coset_def using invx 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: 77362
diff changeset
  1955
        have "(x \<otimes> x' \<otimes> inv x) \<otimes> x = (x \<otimes> x') \<otimes> (inv x \<otimes> 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: 77362
diff changeset
  1956
          by (metis Units_eq Units_inv_Units m_assoc m_closed x'(1) x(1))
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: 77362
diff changeset
  1957
        also have "\<dots> = 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: 77362
diff changeset
  1958
          by (simp add: x 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: 77362
diff changeset
  1959
        finally have "x \<otimes> x' \<otimes> inv x \<otimes> x = 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: 77362
diff changeset
  1960
        with xcoset have "y \<in> {x \<in> carrier G. H #> x \<in> A} #> x" 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: 77362
diff changeset
  1961
      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: 77362
diff changeset
  1962
        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: 77362
diff changeset
  1963
    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: 77362
diff changeset
  1964
  qed 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: 77362
diff changeset
  1965
  with assms 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: 77362
diff changeset
  1966
    by (metis (full_types) factgroup_subgroup_union_char 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: 77362
diff changeset
  1967
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: 77362
diff changeset
  1968
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: 77362
diff changeset
  1969
lemma (in normal) factgroup_subgroup_union_factor:
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: 77362
diff changeset
  1970
  assumes "subgroup A (G 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: 77362
diff changeset
  1971
  shows "A = rcosets\<^bsub>G\<lparr>carrier := \<Union>A\<rparr>\<^esub> 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: 77362
diff changeset
  1972
  using assms subgroup.mem_carrier factgroup_subgroup_union_char by (fastforce simp: RCOSETS_def FactGroup_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: 77362
diff changeset
  1973
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: 77362
diff changeset
  1974
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: 77362
diff changeset
  1975
section  \<open>Flattening the type of group carriers\<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: 77362
diff changeset
  1976
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: 77362
diff changeset
  1977
text \<open>Flattening here means to convert the type of group elements from 'a set to 'a.
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: 77362
diff changeset
  1978
This is possible whenever the empty set is not an element of the group. By Jakob von Raumer\<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: 77362
diff changeset
  1979
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: 77362
diff changeset
  1980
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: 77362
diff changeset
  1981
definition flatten where
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: 77362
diff changeset
  1982
  "flatten (G::('a set, 'b) monoid_scheme) rep = \<lparr>carrier=(rep ` (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: 77362
diff changeset
  1983
      monoid.mult=(\<lambda> x y. rep ((the_inv_into (carrier G) rep x) \<otimes>\<^bsub>G\<^esub> (the_inv_into (carrier G) rep 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: 77362
diff changeset
  1984
      one=rep \<one>\<^bsub>G\<^esub> \<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: 77362
diff changeset
  1985
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: 77362
diff changeset
  1986
lemma flatten_set_group_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: 77362
diff changeset
  1987
  assumes group: "group 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: 77362
diff changeset
  1988
  assumes inj: "inj_on rep (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: 77362
diff changeset
  1989
  shows "rep \<in> hom G (flatten G rep)"
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: 77362
diff changeset
  1990
  by (force simp add: hom_def flatten_def inj the_inv_into_f_f)
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: 77362
diff changeset
  1991
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: 77362
diff changeset
  1992
lemma flatten_set_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: 77362
diff changeset
  1993
  assumes "group G" "inj_on rep (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: 77362
diff changeset
  1994
  shows "group (flatten G rep)"
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: 77362
diff changeset
  1995
proof (rule groupI)
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: 77362
diff changeset
  1996
  fix x 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: 77362
diff changeset
  1997
  assume "x \<in> carrier (flatten G rep)" and "y \<in> carrier (flatten G rep)"
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: 77362
diff changeset
  1998
  then show "x \<otimes>\<^bsub>flatten G rep\<^esub> y \<in> carrier (flatten G rep)" 
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: 77362
diff changeset
  1999
    using assms group.surj_const_mult the_inv_into_f_f by (fastforce simp: flatten_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: 77362
diff changeset
  2000
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: 77362
diff changeset
  2001
  show "\<one>\<^bsub>flatten G rep\<^esub> \<in> carrier (flatten G rep)" 
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: 77362
diff changeset
  2002
    unfolding flatten_def by (simp add: assms group.is_monoid)
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: 77362
diff changeset
  2003
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: 77362
diff changeset
  2004
  fix x y z
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: 77362
diff changeset
  2005
  assume "x \<in> carrier (flatten G rep)" "y \<in> carrier (flatten G rep)" "z \<in> carrier (flatten G rep)"
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: 77362
diff changeset
  2006
  then show "x \<otimes>\<^bsub>flatten G rep\<^esub> y \<otimes>\<^bsub>flatten G rep\<^esub> z = x \<otimes>\<^bsub>flatten G rep\<^esub> (y \<otimes>\<^bsub>flatten G rep\<^esub> z)"
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: 77362
diff changeset
  2007
    by (auto simp: assms flatten_def group.is_monoid monoid.m_assoc monoid.m_closed the_inv_into_f_f)
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: 77362
diff changeset
  2008
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: 77362
diff changeset
  2009
  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: 77362
diff changeset
  2010
  assume x: "x \<in> carrier (flatten G rep)"
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: 77362
diff changeset
  2011
  then show "\<one>\<^bsub>flatten G rep\<^esub> \<otimes>\<^bsub>flatten G rep\<^esub> x = 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: 77362
diff changeset
  2012
    by (auto simp: assms group.is_monoid the_inv_into_f_f flatten_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: 77362
diff changeset
  2013
  then have "\<exists>y\<in>carrier G. rep (y \<otimes>\<^bsub>G\<^esub> z) = rep \<one>\<^bsub>G\<^esub>" if "z \<in> carrier G" for z
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: 77362
diff changeset
  2014
    by (metis \<open>group G\<close> group.l_inv_ex that)
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: 77362
diff changeset
  2015
  with assms x show "\<exists>y\<in>carrier (flatten G rep). y \<otimes>\<^bsub>flatten G rep\<^esub> x = \<one>\<^bsub>flatten G rep\<^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: 77362
diff changeset
  2016
    by (auto simp: flatten_def the_inv_into_f_f)
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: 77362
diff changeset
  2017
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: 77362
diff changeset
  2018
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: 77362
diff changeset
  2019
lemma (in normal) flatten_set_group_mod_inj:
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: 77362
diff changeset
  2020
  shows "inj_on (\<lambda>U. SOME g. g \<in> U) (carrier (G 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: 77362
diff changeset
  2021
proof (rule inj_onI)
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: 77362
diff changeset
  2022
  fix U V
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: 77362
diff changeset
  2023
  assume U: "U \<in> carrier (G Mod H)" and V: "V \<in> carrier (G 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: 77362
diff changeset
  2024
  then obtain g h where g: "U = H #> g" "g \<in> carrier G" and h: "V = H #> h" "h \<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: 77362
diff changeset
  2025
    unfolding FactGroup_def RCOSETS_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: 77362
diff changeset
  2026
  hence notempty: "U \<noteq> {}" "V \<noteq> {}" 
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: 77362
diff changeset
  2027
    by (metis empty_iff is_subgroup rcos_self)+
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: 77362
diff changeset
  2028
  assume "(SOME g. g \<in> U) = (SOME g. g \<in> V)"
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: 77362
diff changeset
  2029
  with notempty have "(SOME g. g \<in> U) \<in> U \<inter> V" 
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: 77362
diff changeset
  2030
    by (metis IntI ex_in_conv someI)
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: 77362
diff changeset
  2031
  thus "U = V" 
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: 77362
diff changeset
  2032
    by (metis Int_iff g h is_subgroup repr_independence)
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: 77362
diff changeset
  2033
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: 77362
diff changeset
  2034
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: 77362
diff changeset
  2035
lemma (in normal) flatten_set_group_mod:
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: 77362
diff changeset
  2036
  shows "group (flatten (G Mod H) (\<lambda>U. SOME g. g \<in> U))"
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: 77362
diff changeset
  2037
  by (simp add: factorgroup_is_group flatten_set_group flatten_set_group_mod_inj)
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: 77362
diff changeset
  2038
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: 77362
diff changeset
  2039
lemma (in normal) flatten_set_group_mod_iso:
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: 77362
diff changeset
  2040
  shows "(\<lambda>U. SOME g. g \<in> U) \<in> iso (G Mod H) (flatten (G Mod H) (\<lambda>U. SOME g. g \<in> U))"
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: 77362
diff changeset
  2041
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: 77362
diff changeset
  2042
  have "(\<lambda>U. SOME g. g \<in> U) \<in> hom (G Mod H) (flatten (G Mod H) (\<lambda>U. SOME g. g \<in> U))"
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: 77362
diff changeset
  2043
    using factorgroup_is_group flatten_set_group_hom flatten_set_group_mod_inj 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: 77362
diff changeset
  2044
  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: 77362
diff changeset
  2045
  have "inj_on (\<lambda>U. SOME g. g \<in> U) (carrier (G 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: 77362
diff changeset
  2046
    using flatten_set_group_mod_inj 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: 77362
diff changeset
  2047
  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: 77362
diff changeset
  2048
    by (simp add: iso_def bij_betw_def flatten_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: 77362
diff changeset
  2049
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: 77362
diff changeset
  2050
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
  2051
end