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