src/HOL/Algebra/Coset.thy
author wenzelm
Wed, 15 Jul 2020 11:56:43 +0200
changeset 72034 452073b64f28
parent 70039 733e256ecdf3
child 73932 fd21b4a93043
permissions -rw-r--r--
clarified signature;
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"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    42
  by (simp add: normal_def normal_axioms_def is_group l_coset_def r_coset_def m_comm subgroup.mem_carrier)
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
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
   471
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   472
lemma (in group) one_is_normal: "{\<one>} \<lhd> G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   473
proof(intro normal_invI)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   474
  show "subgroup {\<one>} G"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   475
    by (simp add: subgroup_def)
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   476
qed simp
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   477
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   478
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   479
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
   480
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   481
lemma (in group) l_repr_independence:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   482
  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
   483
  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
   484
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   485
  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
   486
    using assms(1) unfolding l_coset_def by blast
68604
57721285d4ef elimination of some "smt"
paulson <lp15@cam.ac.uk>
parents: 68582
diff changeset
   487
  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
   488
  proof -
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   489
    have "h' \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   490
      by (meson HG h'(1) subgroup.mem_carrier)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   491
    moreover have "h \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   492
      by (meson HG subgroup.mem_carrier that)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   493
    ultimately show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   494
      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
   495
  qed
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   496
  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
   497
    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
   498
  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
   499
    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
   500
  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
   501
    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
   502
  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
   503
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   504
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   505
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
   506
  "\<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
   507
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
   508
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   509
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
   510
by (force simp add: l_coset_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   511
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   512
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
   513
  "\<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
   514
by (auto simp add: l_coset_def subsetD)
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   515
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   516
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
   517
  "\<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
   518
  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
   519
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   520
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
   521
  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
   522
  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
   523
  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
   524
  unfolding l_coset_def by fastforce
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   525
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   526
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
   527
  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
   528
  shows "H <#> H = H"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   529
proof
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   530
  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
   531
    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
   532
  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
   533
  proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   534
    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
   535
      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
   536
      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
   537
  qed
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   538
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   539
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   540
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   541
subsubsection \<open>Set of Inverses of an \<open>r_coset\<close>.\<close>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   542
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   543
lemma (in normal) rcos_inv:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   544
  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
   545
  shows "set_inv (H #> x) = H #> (inv x)"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   546
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
   547
  fix h
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   548
  assume h: "h \<in> H"
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   549
  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
   550
  proof
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   551
    show "inv x \<otimes> inv h \<otimes> x \<in> H"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   552
      by (simp add: inv_op_closed1 h x)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   553
    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
   554
      by (simp add: h x m_assoc)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   555
  qed
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 14963
diff changeset
   556
  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
   557
  proof
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   558
    show "x \<otimes> inv h \<otimes> inv x \<in> H"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   559
      by (simp add: inv_op_closed2 h x)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   560
    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
   561
      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
   562
  qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   563
qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   564
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   565
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   566
subsubsection \<open>Theorems for \<open><#>\<close> with \<open>#>\<close> or \<open><#\<close>.\<close>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   567
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   568
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
   569
  "\<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
   570
    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
   571
  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
   572
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   573
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
   574
  "\<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
   575
   (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
   576
  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
   577
  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
   578
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   579
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
   580
  "\<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
   581
   (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
   582
  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
   583
                subset l_coset_subset_G rcos_assoc_lcos)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   584
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   585
lemma (in normal) rcos_mult_step2:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   586
     "\<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
   587
      \<Longrightarrow> (H <#> (x <# H)) #> y = (H <#> (H #> x)) #> y"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   588
by (insert coset_eq, simp add: normal_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   589
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   590
lemma (in normal) rcos_mult_step3:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   591
     "\<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
   592
      \<Longrightarrow> (H <#> (H #> x)) #> y = H #> (x \<otimes> y)"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   593
by (simp add: setmult_rcos_assoc coset_mult_assoc
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   594
              subgroup_mult_id normal.axioms subset normal_axioms)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   595
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   596
lemma (in normal) rcos_sum:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   597
     "\<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
   598
      \<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
   599
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
   600
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   601
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
   602
  \<comment> \<open>generalizes \<open>subgroup_mult_id\<close>\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   603
  by (auto simp add: RCOSETS_def subset
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   604
        setmult_rcos_assoc subgroup_mult_id normal.axioms normal_axioms)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   605
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   606
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   607
subsubsection\<open>An Equivalence Relation\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   608
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
   609
definition
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
   610
  r_congruent :: "[('a,'b)monoid_scheme, 'a set] \<Rightarrow> ('a*'a)set"  ("rcong\<index> _")
67091
1393c2340eec more symbols;
wenzelm
parents: 65035
diff changeset
   611
  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
   612
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   613
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   614
lemma (in subgroup) equiv_rcong:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   615
   assumes "group G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   616
   shows "equiv (carrier G) (rcong H)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   617
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   618
  interpret group G by fact
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   619
  show ?thesis
40815
6e2d17cc0d1d equivI has replaced equiv.intro
haftmann
parents: 39910
diff changeset
   620
  proof (intro equivI)
30198
922f944f03b2 name changes
nipkow
parents: 29237
diff changeset
   621
    show "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
   622
      by (auto simp add: r_congruent_def refl_on_def)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   623
  next
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   624
    show "sym (rcong H)"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   625
    proof (simp add: r_congruent_def sym_def, clarify)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   626
      fix x y
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   627
      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
   628
         and "inv x \<otimes> y \<in> H"
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 41528
diff changeset
   629
      hence "inv (inv x \<otimes> y) \<in> H" by simp
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   630
      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
   631
    qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   632
  next
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   633
    show "trans (rcong H)"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   634
    proof (simp add: r_congruent_def trans_def, clarify)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   635
      fix x y z
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   636
      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
   637
         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
   638
      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
   639
      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
   640
        by (simp add: m_assoc del: r_inv Units_r_inv)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   641
      thus "inv x \<otimes> z \<in> H" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   642
    qed
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   643
  qed
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   644
qed
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   645
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   646
text\<open>Equivalence classes of \<open>rcong\<close> correspond to left cosets.
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   647
  Was there a mistake in the definitions? I'd have expected them to
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   648
  correspond to right cosets.\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   649
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   650
(* CB: This is correct, but subtle.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   651
   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
   652
   Jacobson, this is what the majority of group theory literature does.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   653
   He then defines the notion of congruence relation ~ over monoids as
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   654
   equivalence relation with a ~ a' & b ~ b' \<Longrightarrow> a*b ~ a'*b'.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   655
   Our notion of right congruence induced by K: rcong K appears only in
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   656
   the context where K is a normal subgroup.  Jacobson doesn't name it.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   657
   But in this context left and right cosets are identical.
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   658
*)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   659
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   660
lemma (in subgroup) l_coset_eq_rcong:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   661
  assumes "group G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   662
  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
   663
  shows "a <# H = (rcong H) `` {a}"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   664
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   665
  interpret group G by fact
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   666
  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
   667
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   668
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   669
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   670
subsubsection\<open>Two Distinct Right Cosets are Disjoint\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   671
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   672
lemma (in group) rcos_equation:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   673
  assumes "subgroup H G"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   674
  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
   675
  shows "hb \<otimes> a \<in> (\<Union>h\<in>H. {h \<otimes> b})"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   676
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   677
  interpret subgroup H G by fact
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   678
  from p show ?thesis 
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   679
    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
   680
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   681
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   682
lemma (in group) rcos_disjoint:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   683
  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
   684
  shows "pairwise disjnt (rcosets H)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   685
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   686
  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
   687
  show ?thesis
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   688
    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
   689
    by (blast intro: rcos_equation assms sym)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   690
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   691
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   692
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   693
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
   694
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   695
text \<open>The relation is a congruence\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   696
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   697
lemma (in normal) congruent_rcong:
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   698
  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
   699
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
   700
  fix a b c
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   701
  assume abrcong: "(a, b) \<in> rcong H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   702
    and ccarr: "c \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   703
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   704
  from abrcong
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   705
      have acarr: "a \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   706
        and bcarr: "b \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   707
        and abH: "inv a \<otimes> b \<in> H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   708
      unfolding r_congruent_def
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   709
      by fast+
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   710
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   711
  note carr = acarr bcarr ccarr
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   712
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   713
  from ccarr and abH
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   714
      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
   715
  moreover
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   716
      from carr and inv_closed
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   717
      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
   718
      by (force cong: m_assoc)
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   719
  moreover
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   720
      from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   721
      have "\<dots> = (inv (a \<otimes> c)) \<otimes> (b \<otimes> c)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   722
      by (simp add: inv_mult_group)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   723
  ultimately
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   724
      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
   725
  from carr and this
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   726
     have "(b \<otimes> c) \<in> (a \<otimes> c) <# H"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   727
     by (simp add: lcos_module_rev[OF is_group])
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   728
  from carr and this and is_subgroup
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   729
     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
   730
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   731
  fix a b c
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   732
  assume abrcong: "(a, b) \<in> rcong H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   733
    and ccarr: "c \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   734
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 41528
diff changeset
   735
  from ccarr have "c \<in> Units G" by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   736
  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
   737
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   738
  from abrcong
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   739
      have acarr: "a \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   740
       and bcarr: "b \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   741
       and abH: "inv a \<otimes> b \<in> H"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   742
      by (unfold r_congruent_def, fast+)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   743
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   744
  note carr = acarr bcarr ccarr
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   745
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   746
  from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   747
     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
   748
  also from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   749
      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
   750
  also from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   751
      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
   752
  also from carr and inv_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   753
      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
   754
  finally
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   755
      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
   756
  from abH and this
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   757
      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
   758
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   759
  from carr and this
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   760
     have "(c \<otimes> b) \<in> (c \<otimes> a) <# H"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   761
     by (simp add: lcos_module_rev[OF is_group])
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   762
  from carr and this and is_subgroup
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   763
     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
   764
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19931
diff changeset
   765
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   766
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   767
subsection \<open>Order of a Group and Lagrange's Theorem\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   768
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   769
definition
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   770
  order :: "('a, 'b) monoid_scheme \<Rightarrow> nat"
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   771
  where "order S = card (carrier S)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   772
61628
8dd2bd4fe30b add lemmas about monoids and groups
Andreas Lochbihler
parents: 61382
diff changeset
   773
lemma (in monoid) order_gt_0_iff_finite: "0 < order G \<longleftrightarrow> finite (carrier G)"
8dd2bd4fe30b add lemmas about monoids and groups
Andreas Lochbihler
parents: 61382
diff changeset
   774
by(auto simp add: order_def card_gt_0_iff)
8dd2bd4fe30b add lemmas about monoids and groups
Andreas Lochbihler
parents: 61382
diff changeset
   775
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   776
lemma (in group) rcosets_part_G:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   777
  assumes "subgroup H G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   778
  shows "\<Union>(rcosets H) = carrier G"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   779
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   780
  interpret subgroup H G by fact
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   781
  show ?thesis
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   782
    unfolding RCOSETS_def r_coset_def by auto
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   783
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   784
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   785
lemma (in group) cosets_finite:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   786
     "\<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
   787
  unfolding RCOSETS_def
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   788
  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
   789
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62347
diff changeset
   790
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
   791
lemma (in group) inj_on_f:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   792
  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
   793
  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
   794
proof 
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   795
  fix x y
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   796
  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
   797
  then have "x \<in> carrier G" "y \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   798
    using assms r_coset_subset_G by blast+
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   799
  with xy a show "x = y"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   800
    by auto
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   801
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   802
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   803
lemma (in group) inj_on_g:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   804
    "\<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
   805
by (force simp add: inj_on_def subsetD)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   806
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   807
(* ************************************************************************** *)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   808
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   809
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
   810
  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
   811
  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
   812
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   813
  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
   814
    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
   815
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   816
  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
   817
  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
   818
  proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   819
    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
   820
    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
   821
      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
   822
    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
   823
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   824
  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
   825
  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
   826
    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
   827
  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
   828
    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
   829
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   830
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   831
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
   832
  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
   833
  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
   834
  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
   835
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   836
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
   837
  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
   838
  shows "finite R"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   839
  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
   840
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   841
(* ************************************************************************** *)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   842
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   843
lemma (in group) rcosets_subset_PowG:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   844
     "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
   845
  using rcosets_part_G by auto
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   846
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   847
proposition (in group) lagrange_finite:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   848
  assumes "finite(carrier G)" and HG: "subgroup H G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   849
  shows "card(rcosets H) * card(H) = order(G)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   850
proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   851
  have "card H * card (rcosets H) = card (\<Union>(rcosets H))"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   852
  proof (rule card_partition)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   853
    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
   854
      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
   855
  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
   856
  then show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   857
    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
   858
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   859
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   860
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
   861
  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
   862
  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
   863
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
   864
  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
   865
next
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   866
  case False 
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   867
  thus ?thesis
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   868
  proof (cases "finite H")
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   869
    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
   870
  next
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   871
    case True 
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   872
    have "infinite (rcosets H)"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   873
    proof 
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   874
      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
   875
      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
   876
      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
   877
        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
   878
              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
   879
      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
   880
        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
   881
      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
   882
        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
   883
      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
   884
      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
   885
                                rcosets_part_G subgroup.one_closed by fastforce
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   886
      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
   887
    qed
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   888
    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
   889
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   890
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   891
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   892
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
   893
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
   894
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   895
definition
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   896
  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
   897
    \<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
   898
   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
   899
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   900
lemma (in normal) setmult_closed:
65035
b46fe5138cb0 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra
haftmann
parents: 64587
diff changeset
   901
     "\<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
   902
by (auto simp add: rcos_sum RCOSETS_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   903
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   904
lemma (in normal) setinv_closed:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   905
     "K \<in> rcosets H \<Longrightarrow> set_inv K \<in> rcosets H"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   906
by (auto simp add: rcos_inv RCOSETS_def)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   907
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   908
lemma (in normal) rcosets_assoc:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   909
     "\<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
   910
      \<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
   911
  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
   912
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   913
lemma (in subgroup) subgroup_in_rcosets:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26310
diff changeset
   914
  assumes "group G"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   915
  shows "H \<in> rcosets H"
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   916
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   917
  interpret group G by fact
26203
9625f3579b48 explicit referencing of background facts;
wenzelm
parents: 23463
diff changeset
   918
  from _ subgroup_axioms have "H #> \<one> = H"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   919
    by (rule coset_join2) auto
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   920
  then show ?thesis
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   921
    by (auto simp add: RCOSETS_def)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   922
qed
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   923
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   924
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
   925
     "M \<in> rcosets H \<Longrightarrow> set_inv M <#> M = H"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40815
diff changeset
   926
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
   927
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   928
theorem (in normal) factorgroup_is_group:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   929
  "group (G Mod H)"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   930
  unfolding FactGroup_def
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   931
  apply (rule groupI)
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   932
    apply (simp add: setmult_closed)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   933
   apply (simp add: normal_imp_subgroup subgroup_in_rcosets [OF is_group])
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   934
  apply (simp add: restrictI setmult_closed rcosets_assoc)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   935
 apply (simp add: normal_imp_subgroup
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   936
                  subgroup_in_rcosets rcosets_mult_eq)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   937
apply (auto dest: rcosets_inv_mult_group_eq simp add: setinv_closed)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   938
done
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   939
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   940
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
   941
  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
   942
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   943
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
   944
  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
   945
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   946
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
   947
  by (auto simp: FactGroup_def)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
   948
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   949
lemma (in normal) inv_FactGroup:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   950
  assumes "X \<in> carrier (G Mod H)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   951
  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
   952
proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   953
  have X: "X \<in> rcosets H"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   954
    using assms by (simp add: FactGroup_def)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   955
  moreover have "set_inv X <#> X = H"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   956
    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
   957
  moreover have "Group.group (G Mod H)"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   958
    using normal.factorgroup_is_group normal_axioms by blast
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   959
  moreover have "set_inv X \<in> rcosets H"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   960
    by (simp add: \<open>X \<in> rcosets H\<close> setinv_closed)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   961
  ultimately show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   962
    by (simp add: FactGroup_def group.inv_equality)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
   963
qed
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   964
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
   965
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
   966
  \<^term>\<open>G Mod H\<close>\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   967
lemma (in normal) r_coset_hom_Mod:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   968
  "(\<lambda>a. H #> a) \<in> hom G (G Mod H)"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   969
  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
   970
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
   971
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   972
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
   973
  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
   974
  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
   975
  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
   976
  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
   977
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   978
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
   979
  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
   980
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
   981
  have "N \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   982
    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
   983
  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
   984
    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
   985
  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
   986
  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
   987
  then show "x \<otimes>\<^bsub>G Mod N\<^esub> y = y \<otimes>\<^bsub>G Mod N\<^esub> x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   988
    apply (simp add: FactGroup_def subgroup_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   989
    apply (rule set_mult_commute)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   990
    using assms apply (auto simp: subgroup_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   991
    done
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   992
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   993
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   994
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   995
lemma FactGroup_universal:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   996
  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
   997
    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
   998
  obtains g
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   999
  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
  1000
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1001
  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
  1002
    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
  1003
  show thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1004
  proof
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1005
    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
  1006
    proof (rule homI)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1007
      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
  1008
        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
  1009
      proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1010
        from that
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1011
        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
  1012
          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
  1013
        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
  1014
           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
  1015
        then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1016
          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
  1017
      qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1018
    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
  1019
  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
  1020
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1021
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1022
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1023
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
  1024
  fixes k::nat
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1025
  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
  1026
  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
  1027
proof (induction k)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1028
  case 0
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1029
  then show ?case
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1030
    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
  1031
next
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1032
  case (Suc k)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1033
  then show ?case
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1034
    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
  1035
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1036
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1037
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
  1038
  fixes k::int
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1039
  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
  1040
  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
  1041
  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
  1042
         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
  1043
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1044
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1045
subsection\<open>The First Isomorphism Theorem\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1046
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1047
text\<open>The quotient by the kernel of a homomorphism is isomorphic to the
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1048
  range of that homomorphism.\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1049
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
  1050
definition
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
  1051
  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
  1052
    \<comment> \<open>the kernel of a homomorphism\<close>
67091
1393c2340eec more symbols;
wenzelm
parents: 65035
diff changeset
  1053
  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
  1054
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1055
lemma (in group_hom) subgroup_kernel: "subgroup (kernel G H h) G"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1056
  by (auto simp add: kernel_def group.intro is_group intro: subgroup.intro)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1057
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1058
text\<open>The kernel of a homomorphism is a normal subgroup\<close>
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1059
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
  1060
  apply (simp only: G.normal_inv_iff subgroup_kernel)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1061
  apply (simp add: kernel_def)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1062
  done
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1063
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1064
lemma iso_kernel_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1065
  assumes "group G" "group H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1066
  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
  1067
    (is "?lhs = ?rhs")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1068
proof (intro iffI conjI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1069
  assume f: ?lhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1070
  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
  1071
    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
  1072
  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
  1073
    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
  1074
    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
  1075
  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
  1076
    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
  1077
next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1078
  assume ?rhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1079
  with assms show ?lhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1080
    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
  1081
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1082
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1083
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1084
lemma (in group_hom) FactGroup_nonempty:
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1085
  assumes X: "X \<in> carrier (G Mod kernel G H h)"
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1086
  shows "X \<noteq> {}"
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1087
proof -
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1088
  from X
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1089
  obtain g where "g \<in> carrier G"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1090
             and "X = kernel G H h #> g"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1091
    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
  1092
  thus ?thesis
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1093
   by (auto simp add: kernel_def r_coset_def image_def intro: hom_one)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1094
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1095
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1096
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1097
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
  1098
  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
  1099
  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
  1100
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1101
  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
  1102
    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
  1103
  proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1104
    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
  1105
      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
  1106
         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
  1107
    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
  1108
      by blast
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1109
    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
  1110
      by (simp add: that)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1111
    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
  1112
      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
  1113
    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
  1114
    then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1115
      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
  1116
  qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1117
  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
  1118
    by metis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1119
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
  1120
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1121
lemma (in group_hom) FactGroup_the_elem_mem:
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1122
  assumes X: "X \<in> carrier (G Mod (kernel G H h))"
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1123
  shows "the_elem (h`X) \<in> carrier H"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1124
proof -
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1125
  from X
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1126
  obtain g where g: "g \<in> carrier G"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1127
             and "X = kernel G H h #> g"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1128
    by (auto simp add: FactGroup_def RCOSETS_def)
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61628
diff changeset
  1129
  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
  1130
  thus ?thesis by (auto simp add: g)
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1131
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1132
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1133
lemma (in group_hom) FactGroup_hom:
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1134
     "(\<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
  1135
proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1136
  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
  1137
    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
  1138
  proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1139
    obtain g and g'
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1140
      where "g \<in> carrier G" and "g' \<in> carrier G"
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1141
        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
  1142
      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
  1143
    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
  1144
      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
  1145
      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
  1146
    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
  1147
      by (auto dest!: FactGroup_nonempty intro!: image_eqI
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1148
          simp add: set_mult_def
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1149
          subsetD [OF Xsub] subsetD [OF X'sub])
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1150
    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
  1151
      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
  1152
  qed
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1153
  then show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1154
    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
  1155
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1156
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1157
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 57512
diff changeset
  1158
text\<open>Lemma for the following injectivity result\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1159
lemma (in group_hom) FactGroup_subset:
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1160
  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
  1161
  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
  1162
  unfolding kernel_def r_coset_def
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1163
proof clarsimp
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1164
  fix y 
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1165
  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
  1166
  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
  1167
    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
  1168
qed
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1169
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1170
lemma (in group_hom) FactGroup_inj_on:
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1171
     "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
  1172
proof (simp add: inj_on_def, clarify)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1173
  fix X and X'
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1174
  assume X:  "X  \<in> carrier (G Mod kernel G H h)"
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1175
     and X': "X' \<in> carrier (G Mod kernel G H h)"
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1176
  then
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1177
  obtain g and g'
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1178
           where gX: "g \<in> carrier G"  "g' \<in> carrier G"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1179
              "X = kernel G H h #> g" "X' = kernel G H h #> g'"
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
  1180
    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
  1181
  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
  1182
    by (force simp add: kernel_def r_coset_def image_def)+
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1183
  assume "the_elem (h ` X) = the_elem (h ` X')"
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1184
  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
  1185
    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
  1186
  show "X=X'" by (rule equalityI) (simp_all add: FactGroup_subset h gX)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1187
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1188
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
  1189
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
  1190
homomorphism from the quotient group\<close>
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1191
lemma (in group_hom) FactGroup_onto:
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1192
  assumes h: "h ` carrier G = carrier H"
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1193
  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
  1194
proof
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1195
  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
  1196
    by (auto simp add: FactGroup_the_elem_mem)
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 35849
diff changeset
  1197
  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
  1198
  proof
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1199
    fix y
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1200
    assume y: "y \<in> carrier H"
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1201
    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
  1202
      by (blast elim: equalityE)
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1203
    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
  1204
      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
  1205
    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
  1206
      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
  1207
      apply (subst the_elem_image_unique)
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61628
diff changeset
  1208
      apply auto
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61628
diff changeset
  1209
      done
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1210
  qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1211
qed
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1212
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1213
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
  1214
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
  1215
 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
  1216
theorem (in group_hom) FactGroup_iso_set:
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1217
  "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
  1218
   \<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
  1219
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
  1220
              FactGroup_onto)
14803
f7557773cc87 more group isomorphisms
paulson
parents: 14761
diff changeset
  1221
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1222
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
  1223
  "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
  1224
   \<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
  1225
  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
  1226
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1227
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1228
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
  1229
  "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
  1230
  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
  1231
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1232
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
  1233
  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
  1234
  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
  1235
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
  1236
  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
  1237
  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
  1238
  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
  1239
    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
  1240
  thus "g1 = g2"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1241
    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
  1242
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1243
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1244
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
  1245
  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
  1246
proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1247
  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
  1248
    unfolding kernel_def
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1249
  proof (auto)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1250
    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
  1251
      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
  1252
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1253
next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1254
  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
  1255
    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
  1256
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1257
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1258
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
  1259
  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
  1260
proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1261
  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
  1262
    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
  1263
  thus ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1264
    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
  1265
    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
  1266
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1267
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1268
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
  1269
  assumes "subgroup I G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1270
  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
  1271
  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
  1272
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1273
lemma set_mult_hom:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1274
  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
  1275
  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
  1276
proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1277
  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
  1278
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1279
    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
  1280
    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
  1281
      unfolding set_mult_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1282
    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
  1283
      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
  1284
    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
  1285
      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
  1286
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1287
next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1288
  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
  1289
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1290
    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
  1291
    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
  1292
      unfolding set_mult_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1293
    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
  1294
      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
  1295
    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
  1296
      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
  1297
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1298
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1299
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1300
corollary coset_hom:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1301
  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
  1302
  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
  1303
  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
  1304
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1305
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
  1306
  assumes "I \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1307
  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
  1308
proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1309
  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
  1310
    unfolding kernel_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1311
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1312
  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
  1313
    unfolding kernel_def by force
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1314
  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
  1315
    using assms by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1316
  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
  1317
    unfolding set_mult_def by force+
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1318
  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
  1319
    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
  1320
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68975
diff changeset
  1321
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1322
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
  1323
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1324
definition trivial_homomorphism where
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1325
 "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
  1326
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1327
lemma trivial_homomorphism_kernel:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1328
   "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
  1329
  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
  1330
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1331
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
  1332
   "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
  1333
  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
  1334
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1335
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1336
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
  1337
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1338
lemma group_Int_image_ker:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1339
  assumes f: "f \<in> hom G H" and g: "g \<in> hom H K" and "inj_on (g \<circ> f) (carrier G)" "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
  1340
  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
  1341
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1342
  have "(f ` carrier G) \<inter> (kernel H K g) \<subseteq> {\<one>\<^bsub>H\<^esub>}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1343
    using assms
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1344
    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
  1345
    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
  1346
  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
  1347
    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
  1348
  moreover have "one H \<in> kernel H K g"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1349
    apply (simp add: kernel_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1350
    using g group.is_monoid hom_one \<open>group H\<close> \<open>group K\<close> by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1351
  ultimately show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1352
    by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1353
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1354
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1355
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1356
lemma group_sum_image_ker:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1357
  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
  1358
     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
  1359
  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
  1360
proof
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1361
  show "?lhs \<subseteq> ?rhs"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1362
    apply (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
  1363
    by (meson Group.group_def assms(5) f hom_carrier image_eqI monoid.m_closed subset_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1364
  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
  1365
    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
  1366
  proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1367
    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
  1368
      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
  1369
    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
  1370
      by (metis image_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1371
    with assms have "inv\<^bsub>H\<^esub> f x \<otimes>\<^bsub>H\<^esub> y \<in> carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1372
      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
  1373
    moreover
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1374
    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
  1375
    proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1376
      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
  1377
        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
  1378
      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
  1379
        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
  1380
      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
  1381
        using assms x(1)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1382
        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
  1383
      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
  1384
        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
  1385
      finally show ?thesis .
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1386
    qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1387
    moreover
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1388
    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
  1389
      using x y
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1390
      by (metis (no_types, hide_lams) assms(5) f group.inv_solve_left group.subgroup_self hom_carrier image_subset_iff subgroup_def that)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1391
    ultimately
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1392
    show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1393
      using x y by force
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1394
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1395
  then show "?rhs \<subseteq> ?lhs"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1396
    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
  1397
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1398
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1399
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1400
lemma group_sum_ker_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1401
  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
  1402
     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
  1403
   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
  1404
proof
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1405
  show "?lhs \<subseteq> ?rhs"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1406
    apply (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
  1407
    by (meson Group.group_def \<open>group H\<close> f hom_carrier image_eqI monoid.m_closed subset_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1408
  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
  1409
    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
  1410
  proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1411
    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
  1412
      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
  1413
    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
  1414
      by (metis image_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1415
    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
  1416
      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
  1417
    moreover
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1418
    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
  1419
    proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1420
      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
  1421
        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
  1422
      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
  1423
        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
  1424
      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
  1425
        using assms x(1)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1426
        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
  1427
      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
  1428
        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
  1429
        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
  1430
      finally show ?thesis .
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1431
    qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1432
    moreover
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1433
    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
  1434
      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
  1435
    ultimately
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1436
    show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1437
      using x y by force
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1438
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1439
  then show "?rhs \<subseteq> ?lhs"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1440
    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
  1441
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1442
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1443
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
  1444
  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
  1445
  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
  1446
        "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
  1447
  using assms
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1448
  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
  1449
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1450
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
  1451
  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
  1452
     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
  1453
   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
  1454
          "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
  1455
  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
  1456
  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
  1457
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1458
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1459
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
  1460
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
  1461
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1462
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
  1463
  assumes "group K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1464
    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
  1465
    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
  1466
  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
  1467
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
  1468
  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
  1469
    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
  1470
         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
  1471
  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
  1472
  proof-
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1473
    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
  1474
    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
  1475
    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
  1476
      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
  1477
    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
  1478
      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
  1479
    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
  1480
      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
  1481
    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
  1482
      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
  1483
    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
  1484
      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
  1485
    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
  1486
      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
  1487
    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
  1488
    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
  1489
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1490
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1491
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1492
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
  1493
  assumes "group K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1494
    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
  1495
    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
  1496
  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
  1497
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1498
proof-
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1499
  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
  1500
    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
  1501
  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
  1502
                \<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
  1503
    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
  1504
  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
  1505
                 \<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
  1506
    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
  1507
    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
  1508
  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
  1509
                                     carrier (G \<times>\<times> K Mod H \<times> N)"
68687
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1510
  proof -
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1511
    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
  1512
      using R by force
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1513
    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
  1514
      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
  1515
    show ?thesis
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1516
      unfolding image_def by (auto simp: intro: 1 2)
2976a4a3b126 de-applying and simplification
paulson <lp15@cam.ac.uk>
parents: 68604
diff changeset
  1517
  qed
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1518
  ultimately show ?thesis
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1519
    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
  1520
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1521
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1522
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
  1523
  assumes "group K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1524
    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
  1525
    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
  1526
  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
  1527
  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
  1528
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69749
diff changeset
  1529
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
  1530
  assumes "group K"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1531
    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
  1532
    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
  1533
  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
  1534
  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
  1535
        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
  1536
  by blast
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1537
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
  1538
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
  1539
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
  1540
(*A group multiplied by a subgroup stays the same*)
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1541
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
  1542
  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
  1543
  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
  1544
proof
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1545
  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
  1546
    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
  1547
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
  1548
  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
  1549
  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
  1550
    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
  1551
  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
  1552
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
  1553
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
  1554
(*Same lemma as above, but everything is included in a subgroup*)
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1555
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
  1556
  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
  1557
  shows "H <#> N = H"
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1558
  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
  1559
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
  1560
(*A normal subgroup is commutative with set_mult*)
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1561
lemma (in group) commut_normal:
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1562
  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
  1563
  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
  1564
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
  1565
  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
  1566
  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
  1567
  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
  1568
  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
  1569
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
  1570
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
  1571
(*Same lemma as above, but everything is included in a subgroup*)
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1572
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
  1573
  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
  1574
    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
  1575
  shows "K <#> N = N <#> K"
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1576
  using group.commut_normal[OF subgroup.subgroup_is_group[OF assms(1) group_axioms] assms(3,2)] by simp
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1577
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
  1578
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
  1579
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
  1580
subsubsection "Lemmas about intersection and normal subgroups"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
  1581
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
  1582
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
  1583
  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
  1584
    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
  1585
    and "H1\<lhd>G\<lparr>carrier := H\<rparr>"
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1586
  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
  1587
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
  1588
  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
  1589
    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
  1590
  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
  1591
  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
  1592
    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
  1593
      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
  1594
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
  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
  1596
    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
  1597
    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
  1598
      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
  1599
        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
  1600
    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
  1601
      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
  1602
    next
68555
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1603
      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
  1604
        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
  1605
      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
  1606
      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
  1607
        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
  1608
    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
  1609
    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
  1610
  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
  1611
    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
  1612
    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
  1613
      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
  1614
        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
  1615
      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
  1616
        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
  1617
      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
  1618
      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
  1619
      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
  1620
        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
  1621
      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
  1622
      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
  1623
        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
  1624
      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
  1625
        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
  1626
      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
  1627
        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
  1628
      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
  1629
        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
  1630
      have xH:"x\<in>carrier(GH)"
22d51874f37d a few more lemmas from Paulo and Martin
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1631
        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
  1632
      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
  1633
        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
  1634
      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
  1635
        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
  1636
      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
  1637
        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
  1638
      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
  1639
        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
  1640
        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
  1641
      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
  1642
      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
  1643
      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
  1644
    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
  1645
  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
  1646
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
  1647
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69895
diff changeset
  1648
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
  1649
  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
  1650
    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
  1651
  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
  1652
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
  1653
  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
  1654
  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
  1655
  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
  1656
  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
  1657
  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
  1658
    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
  1659
  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
  1660
  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
  1661
 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
  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
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
  1664
end