src/HOL/Algebra/Coset.thy
author obua
Sun, 09 May 2004 23:04:36 +0200
changeset 14722 8e739a6eaf11
parent 14706 71590b7733b7
child 14747 2eaff69d3d10
permissions -rw-r--r--
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
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
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     2
    ID:         $Id$
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     3
    Author:     Florian Kammueller, with new proofs by L C Paulson
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
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     6
header{*Theory of Cosets*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     7
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     8
theory Coset = Group + Exponent:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     9
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    10
declare (in group) l_inv [simp] and r_inv [simp]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    11
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14530
diff changeset
    12
constdefs (structure G)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    13
  r_coset    :: "[_,'a set, 'a] => 'a set"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    14
  "r_coset G H a == (% x. x \<otimes> a) ` H"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    15
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14530
diff changeset
    16
  l_coset    :: "[_, 'a, 'a set] => 'a set"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    17
  "l_coset G a H == (% x. a \<otimes> x) ` H"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    18
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14530
diff changeset
    19
  rcosets  :: "[_, 'a set] => ('a set)set"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    20
  "rcosets G H == r_coset G H ` (carrier G)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    21
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14530
diff changeset
    22
  set_mult  :: "[_, 'a set ,'a set] => 'a set"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    23
  "set_mult G H K == (%(x,y). x \<otimes> y) ` (H \<times> K)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    24
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14530
diff changeset
    25
  set_inv   :: "[_,'a set] => 'a set"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    26
  "set_inv G H == m_inv G ` H"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    27
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14530
diff changeset
    28
  normal     :: "['a set, _] => bool"       (infixl "<|" 60)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    29
  "normal H G == subgroup H G &
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    30
                  (\<forall>x \<in> carrier G. r_coset G H x = l_coset G x H)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    31
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    32
syntax (xsymbols)
13936
d3671b878828 Greatly extended CRing. Added Module.
ballarin
parents: 13889
diff changeset
    33
  normal  :: "['a set, ('a,'b) monoid_scheme] => bool" (infixl "\<lhd>" 60)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    34
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    35
locale coset = group G +
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    36
  fixes rcos      :: "['a set, 'a] => 'a set"     (infixl "#>" 60)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    37
    and lcos      :: "['a, 'a set] => 'a set"     (infixl "<#" 60)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    38
    and setmult   :: "['a set, 'a set] => 'a set" (infixl "<#>" 60)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    39
  defines rcos_def: "H #> x == r_coset G H x"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    40
      and lcos_def: "x <# H == l_coset G x H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    41
      and setmult_def: "H <#> K == set_mult G H K"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    42
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
    43
text {*
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
    44
  Locale @{term coset} provides only syntax.
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
    45
  Logically, groups are cosets.
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
    46
*}
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
    47
lemma (in group) is_coset:
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
    48
  "coset G"
14254
342634f38451 Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
ballarin
parents: 13943
diff changeset
    49
  by (rule coset.intro)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    50
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    51
text{*Lemmas useful for Sylow's Theorem*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    52
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    53
lemma card_inj:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    54
     "[|f \<in> A\<rightarrow>B; inj_on f A; finite A; finite B|] ==> card(A) <= card(B)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    55
apply (rule card_inj_on_le)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    56
apply (auto simp add: Pi_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    57
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    58
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    59
lemma card_bij:
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    60
     "[|f \<in> A\<rightarrow>B; inj_on f A;
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    61
        g \<in> B\<rightarrow>A; inj_on g B; finite A; finite B|] ==> card(A) = card(B)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    62
by (blast intro: card_inj order_antisym)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    63
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    64
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    65
subsection {*Lemmas Using *}
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    66
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    67
lemma (in coset) r_coset_eq: "H #> a = {b . \<exists>h\<in>H. h \<otimes> a = b}"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    68
by (auto simp add: rcos_def r_coset_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    69
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    70
lemma (in coset) l_coset_eq: "a <# H = {b . \<exists>h\<in>H. a \<otimes> h = b}"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    71
by (auto simp add: lcos_def l_coset_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    72
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    73
lemma (in coset) setrcos_eq: "rcosets G H = {C . \<exists>a\<in>carrier G. C = H #> a}"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    74
by (auto simp add: rcosets_def rcos_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    75
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    76
lemma (in coset) set_mult_eq: "H <#> K = {c . \<exists>h\<in>H. \<exists>k\<in>K. c = h \<otimes> k}"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    77
by (simp add: setmult_def set_mult_def image_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    78
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    79
lemma (in coset) coset_mult_assoc:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    80
     "[| M <= carrier G; g \<in> carrier G; h \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    81
      ==> (M #> g) #> h = M #> (g \<otimes> h)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    82
by (force simp add: r_coset_eq m_assoc)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    83
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    84
lemma (in coset) coset_mult_one [simp]: "M <= carrier G ==> M #> \<one> = M"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    85
by (force simp add: r_coset_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    86
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    87
lemma (in coset) coset_mult_inv1:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    88
     "[| M #> (x \<otimes> (inv y)) = M;  x \<in> carrier G ; y \<in> carrier G;
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    89
         M <= carrier G |] ==> M #> x = M #> y"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    90
apply (erule subst [of concl: "%z. M #> x = z #> y"])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    91
apply (simp add: coset_mult_assoc m_assoc)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    92
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    93
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    94
lemma (in coset) coset_mult_inv2:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    95
     "[| M #> x = M #> y;  x \<in> carrier G;  y \<in> carrier G;  M <= carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    96
      ==> M #> (x \<otimes> (inv y)) = M "
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    97
apply (simp add: coset_mult_assoc [symmetric])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    98
apply (simp add: coset_mult_assoc)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    99
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   100
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   101
lemma (in coset) coset_join1:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   102
     "[| H #> x = H;  x \<in> carrier G;  subgroup H G |] ==> x\<in>H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   103
apply (erule subst)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   104
apply (simp add: r_coset_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   105
apply (blast intro: l_one subgroup.one_closed)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   106
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   107
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   108
text{*Locales don't currently work with @{text rule_tac}, so we
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   109
must prove this lemma separately.*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   110
lemma (in coset) solve_equation:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   111
    "\<lbrakk>subgroup H G; x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> \<exists>h\<in>H. h \<otimes> x = y"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   112
apply (rule bexI [of _ "y \<otimes> (inv x)"])
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   113
apply (auto simp add: subgroup.m_closed subgroup.m_inv_closed m_assoc
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   114
                      subgroup.subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   115
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   116
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   117
lemma (in coset) coset_join2:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   118
     "[| x \<in> carrier G;  subgroup H G;  x\<in>H |] ==> H #> x = H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   119
by (force simp add: subgroup.m_closed r_coset_eq solve_equation)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   120
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   121
lemma (in coset) r_coset_subset_G:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   122
     "[| H <= carrier G; x \<in> carrier G |] ==> H #> x <= carrier G"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   123
by (auto simp add: r_coset_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   124
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   125
lemma (in coset) rcosI:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   126
     "[| h \<in> H; H <= carrier G; x \<in> carrier G|] ==> h \<otimes> x \<in> H #> x"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   127
by (auto simp add: r_coset_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   128
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   129
lemma (in coset) setrcosI:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   130
     "[| H <= carrier G; x \<in> carrier G |] ==> H #> x \<in> rcosets G H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   131
by (auto simp add: setrcos_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   132
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   133
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   134
text{*Really needed?*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   135
lemma (in coset) transpose_inv:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   136
     "[| x \<otimes> y = z;  x \<in> carrier G;  y \<in> carrier G;  z \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   137
      ==> (inv x) \<otimes> z = y"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   138
by (force simp add: m_assoc [symmetric])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   139
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   140
lemma (in coset) repr_independence:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   141
     "[| y \<in> H #> x;  x \<in> carrier G; subgroup H G |] ==> H #> x = H #> y"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   142
by (auto simp add: r_coset_eq m_assoc [symmetric]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   143
                   subgroup.subset [THEN subsetD]
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   144
                   subgroup.m_closed solve_equation)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   145
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   146
lemma (in coset) rcos_self: "[| x \<in> carrier G; subgroup H G |] ==> x \<in> H #> x"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   147
apply (simp add: r_coset_eq)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   148
apply (blast intro: l_one subgroup.subset [THEN subsetD]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   149
                    subgroup.one_closed)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   150
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   151
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   152
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   153
subsection {* Normal subgroups *}
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   154
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   155
(*????????????????
13936
d3671b878828 Greatly extended CRing. Added Module.
ballarin
parents: 13889
diff changeset
   156
text "Allows use of theorems proved in the locale coset"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   157
lemma subgroup_imp_coset: "subgroup H G ==> coset G"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   158
apply (drule subgroup_imp_group)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   159
apply (simp add: group_def coset_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   160
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   161
*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   162
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   163
lemma normal_imp_subgroup: "H <| G ==> subgroup H G"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   164
by (simp add: normal_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   165
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   166
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   167
(*????????????????
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   168
lemmas normal_imp_group = normal_imp_subgroup [THEN subgroup_imp_group]
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   169
lemmas normal_imp_coset = normal_imp_subgroup [THEN subgroup_imp_coset]
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   170
*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   171
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   172
lemma (in coset) normal_iff:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   173
     "(H <| G) = (subgroup H G & (\<forall>x \<in> carrier G. H #> x = x <# H))"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   174
by (simp add: lcos_def rcos_def normal_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   175
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   176
lemma (in coset) normal_imp_rcos_eq_lcos:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   177
     "[| H <| G; x \<in> carrier G |] ==> H #> x = x <# H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   178
by (simp add: lcos_def rcos_def normal_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   179
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   180
lemma (in coset) normal_inv_op_closed1:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   181
     "\<lbrakk>H \<lhd> G; x \<in> carrier G; h \<in> H\<rbrakk> \<Longrightarrow> (inv x) \<otimes> h \<otimes> x \<in> H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   182
apply (auto simp add: l_coset_eq r_coset_eq normal_iff)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   183
apply (drule bspec, assumption)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   184
apply (drule equalityD1 [THEN subsetD], blast, clarify)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   185
apply (simp add: m_assoc subgroup.subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   186
apply (erule subst)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   187
apply (simp add: m_assoc [symmetric] subgroup.subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   188
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   189
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   190
lemma (in coset) normal_inv_op_closed2:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   191
     "\<lbrakk>H \<lhd> G; x \<in> carrier G; h \<in> H\<rbrakk> \<Longrightarrow> x \<otimes> h \<otimes> (inv x) \<in> H"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   192
apply (drule normal_inv_op_closed1 [of H "(inv x)"])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   193
apply auto
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   194
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   195
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   196
lemma (in coset) lcos_m_assoc:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   197
     "[| M <= carrier G; g \<in> carrier G; h \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   198
      ==> g <# (h <# M) = (g \<otimes> h) <# M"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   199
by (force simp add: l_coset_eq m_assoc)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   200
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   201
lemma (in coset) lcos_mult_one: "M <= carrier G ==> \<one> <# M = M"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   202
by (force simp add: l_coset_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   203
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   204
lemma (in coset) l_coset_subset_G:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   205
     "[| H <= carrier G; x \<in> carrier G |] ==> x <# H <= carrier G"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   206
by (auto simp add: l_coset_eq subsetD)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   207
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   208
lemma (in coset) l_coset_swap:
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   209
     "[| y \<in> x <# H;  x \<in> carrier G;  subgroup H G |] ==> x \<in> y <# H"
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   210
proof (simp add: l_coset_eq)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   211
  assume "\<exists>h\<in>H. x \<otimes> h = y"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   212
    and x: "x \<in> carrier G"
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   213
    and sb: "subgroup H G"
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   214
  then obtain h' where h': "h' \<in> H & x \<otimes> h' = y" by blast
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   215
  show "\<exists>h\<in>H. y \<otimes> h = x"
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   216
  proof
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   217
    show "y \<otimes> inv h' = x" using h' x sb
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   218
      by (auto simp add: m_assoc subgroup.subset [THEN subsetD])
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   219
    show "inv h' \<in> H" using h' sb
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   220
      by (auto simp add: subgroup.subset [THEN subsetD] subgroup.m_inv_closed)
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   221
  qed
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   222
qed
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   223
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   224
lemma (in coset) l_coset_carrier:
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   225
     "[| y \<in> x <# H;  x \<in> carrier G;  subgroup H G |] ==> y \<in> carrier G"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   226
by (auto simp add: l_coset_eq m_assoc
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   227
                   subgroup.subset [THEN subsetD] subgroup.m_closed)
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   228
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   229
lemma (in coset) l_repr_imp_subset:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   230
  assumes y: "y \<in> x <# H" and x: "x \<in> carrier G" and sb: "subgroup H G"
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   231
  shows "y <# H \<subseteq> x <# H"
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   232
proof -
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   233
  from y
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   234
  obtain h' where "h' \<in> H" "x \<otimes> h' = y" by (auto simp add: l_coset_eq)
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   235
  thus ?thesis using x sb
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   236
    by (auto simp add: l_coset_eq m_assoc
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   237
                       subgroup.subset [THEN subsetD] subgroup.m_closed)
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   238
qed
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   239
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   240
lemma (in coset) l_repr_independence:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   241
  assumes y: "y \<in> x <# H" and x: "x \<in> carrier G" and sb: "subgroup H G"
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   242
  shows "x <# H = y <# H"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   243
proof
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   244
  show "x <# H \<subseteq> y <# H"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   245
    by (rule l_repr_imp_subset,
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   246
        (blast intro: l_coset_swap l_coset_carrier y x sb)+)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   247
  show "y <# H \<subseteq> x <# H" by (rule l_repr_imp_subset [OF y x sb])
14530
e94fd774ecf5 some (much longer) structured proofs
paulson
parents: 14254
diff changeset
   248
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   249
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   250
lemma (in coset) setmult_subset_G:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   251
     "[| H <= carrier G; K <= carrier G |] ==> H <#> K <= carrier G"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   252
by (auto simp add: set_mult_eq subsetD)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   253
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   254
lemma (in coset) subgroup_mult_id: "subgroup H G ==> H <#> H = H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   255
apply (auto simp add: subgroup.m_closed set_mult_eq Sigma_def image_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   256
apply (rule_tac x = x in bexI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   257
apply (rule bexI [of _ "\<one>"])
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   258
apply (auto simp add: subgroup.m_closed subgroup.one_closed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   259
                      r_one subgroup.subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   260
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   261
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   262
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   263
text {* Set of inverses of an @{text r_coset}. *}
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   264
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   265
lemma (in coset) rcos_inv:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   266
  assumes normalHG: "H <| G"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   267
      and xinG:     "x \<in> carrier G"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   268
  shows "set_inv G (H #> x) = H #> (inv x)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   269
proof -
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   270
  have H_subset: "H <= carrier G"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   271
    by (rule subgroup.subset [OF normal_imp_subgroup, OF normalHG])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   272
  show ?thesis
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   273
  proof (auto simp add: r_coset_eq image_def set_inv_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   274
    fix h
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   275
    assume "h \<in> H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   276
      hence "((inv x) \<otimes> (inv h) \<otimes> x) \<otimes> inv x = inv (h \<otimes> x)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   277
        by (simp add: xinG m_assoc inv_mult_group H_subset [THEN subsetD])
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   278
      thus "\<exists>j\<in>H. j \<otimes> inv x = inv (h \<otimes> x)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   279
       using prems
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   280
        by (blast intro: normal_inv_op_closed1 normal_imp_subgroup
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   281
                         subgroup.m_inv_closed)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   282
  next
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   283
    fix h
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   284
    assume "h \<in> H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   285
      hence eq: "(x \<otimes> (inv h) \<otimes> (inv x)) \<otimes> x = x \<otimes> inv h"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   286
        by (simp add: xinG m_assoc H_subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   287
      hence "(\<exists>j\<in>H. j \<otimes> x = inv  (h \<otimes> (inv x))) \<and> h \<otimes> inv x = inv (inv (h \<otimes> (inv x)))"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   288
       using prems
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   289
        by (simp add: m_assoc inv_mult_group H_subset [THEN subsetD],
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   290
            blast intro: eq normal_inv_op_closed2 normal_imp_subgroup
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   291
                         subgroup.m_inv_closed)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   292
      thus "\<exists>y. (\<exists>h\<in>H. h \<otimes> x = y) \<and> h \<otimes> inv x = inv y" ..
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   293
  qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   294
qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   295
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   296
(*The old proof is something like this, but the rule_tac calls make
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   297
illegal references to implicit structures.
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   298
lemma (in coset) rcos_inv:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   299
     "[| H <| G; x \<in> carrier G |] ==> set_inv G (H #> x) = H #> (inv x)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   300
apply (frule normal_imp_subgroup)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   301
apply (auto simp add: r_coset_eq image_def set_inv_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   302
apply (rule_tac x = "(inv x) \<otimes> (inv h) \<otimes> x" in bexI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   303
apply (simp add: m_assoc inv_mult_group subgroup.subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   304
apply (simp add: subgroup.m_inv_closed, assumption+)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   305
apply (rule_tac x = "inv  (h \<otimes> (inv x)) " in exI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   306
apply (simp add: inv_mult_group subgroup.subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   307
apply (rule_tac x = "x \<otimes> (inv h) \<otimes> (inv x)" in bexI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   308
apply (simp add: m_assoc subgroup.subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   309
apply (simp add: normal_inv_op_closed2 subgroup.m_inv_closed)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   310
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   311
*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   312
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   313
lemma (in coset) rcos_inv2:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   314
     "[| H <| G; K \<in> rcosets G H; x \<in> K |] ==> set_inv G K = H #> (inv x)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   315
apply (simp add: setrcos_eq, clarify)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   316
apply (subgoal_tac "x : carrier G")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   317
 prefer 2
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   318
 apply (blast dest: r_coset_subset_G subgroup.subset normal_imp_subgroup)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   319
apply (drule repr_independence)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   320
  apply assumption
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   321
 apply (erule normal_imp_subgroup)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   322
apply (simp add: rcos_inv)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   323
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   324
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   325
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   326
text {* Some rules for @{text "<#>"} with @{text "#>"} or @{text "<#"}. *}
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   327
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   328
lemma (in coset) setmult_rcos_assoc:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   329
     "[| H <= carrier G; K <= carrier G; x \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   330
      ==> H <#> (K #> x) = (H <#> K) #> x"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   331
apply (auto simp add: rcos_def r_coset_def setmult_def set_mult_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   332
apply (force simp add: m_assoc)+
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   333
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   334
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   335
lemma (in coset) rcos_assoc_lcos:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   336
     "[| H <= carrier G; K <= carrier G; x \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   337
      ==> (H #> x) <#> K = H <#> (x <# K)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   338
apply (auto simp add: rcos_def r_coset_def lcos_def l_coset_def
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   339
                      setmult_def set_mult_def Sigma_def image_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   340
apply (force intro!: exI bexI simp add: m_assoc)+
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   341
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   342
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   343
lemma (in coset) rcos_mult_step1:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   344
     "[| H <| G; x \<in> carrier G; y \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   345
      ==> (H #> x) <#> (H #> y) = (H <#> (x <# H)) #> y"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   346
by (simp add: setmult_rcos_assoc normal_imp_subgroup [THEN subgroup.subset]
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   347
              r_coset_subset_G l_coset_subset_G rcos_assoc_lcos)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   348
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   349
lemma (in coset) rcos_mult_step2:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   350
     "[| H <| G; x \<in> carrier G; y \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   351
      ==> (H <#> (x <# H)) #> y = (H <#> (H #> x)) #> y"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   352
by (simp add: normal_imp_rcos_eq_lcos)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   353
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   354
lemma (in coset) rcos_mult_step3:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   355
     "[| H <| G; x \<in> carrier G; y \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   356
      ==> (H <#> (H #> x)) #> y = H #> (x \<otimes> y)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   357
by (simp add: setmult_rcos_assoc r_coset_subset_G coset_mult_assoc
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   358
              setmult_subset_G  subgroup_mult_id
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   359
              subgroup.subset normal_imp_subgroup)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   360
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   361
lemma (in coset) rcos_sum:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   362
     "[| H <| G; x \<in> carrier G; y \<in> carrier G |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   363
      ==> (H #> x) <#> (H #> y) = H #> (x \<otimes> y)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   364
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
   365
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   366
lemma (in coset) setrcos_mult_eq: "[|H <| G; M \<in> rcosets G H|] ==> H <#> M = M"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   367
  -- {* generalizes @{text subgroup_mult_id} *}
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   368
  by (auto simp add: setrcos_eq normal_imp_subgroup subgroup.subset
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   369
    setmult_rcos_assoc subgroup_mult_id)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   370
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   371
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   372
subsection {*Lemmas Leading to Lagrange's Theorem *}
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   373
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   374
lemma (in coset) setrcos_part_G: "subgroup H G ==> \<Union>rcosets G H = carrier G"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   375
apply (rule equalityI)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   376
apply (force simp add: subgroup.subset [THEN subsetD]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   377
                       setrcos_eq r_coset_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   378
apply (auto simp add: setrcos_eq subgroup.subset rcos_self)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   379
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   380
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   381
lemma (in coset) cosets_finite:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   382
     "[| c \<in> rcosets G H;  H <= carrier G;  finite (carrier G) |] ==> finite c"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   383
apply (auto simp add: setrcos_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   384
apply (simp (no_asm_simp) add: r_coset_subset_G [THEN finite_subset])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   385
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   386
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   387
text{*The next two lemmas support the proof of @{text card_cosets_equal},
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   388
since we can't use @{text rule_tac} with explicit terms for @{term f} and
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   389
@{term g}.*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   390
lemma (in coset) inj_on_f:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   391
    "[|H \<subseteq> carrier G;  a \<in> carrier G|] ==> inj_on (\<lambda>y. y \<otimes> inv a) (H #> a)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   392
apply (rule inj_onI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   393
apply (subgoal_tac "x \<in> carrier G & y \<in> carrier G")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   394
 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
   395
apply (simp add: subsetD)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   396
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   397
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   398
lemma (in coset) inj_on_g:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   399
    "[|H \<subseteq> carrier G;  a \<in> carrier G|] ==> inj_on (\<lambda>y. y \<otimes> a) H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   400
by (force simp add: inj_on_def subsetD)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   401
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   402
lemma (in coset) card_cosets_equal:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   403
     "[| c \<in> rcosets G H;  H <= carrier G; finite(carrier G) |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   404
      ==> card c = card H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   405
apply (auto simp add: setrcos_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   406
apply (rule card_bij_eq)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   407
     apply (rule inj_on_f, assumption+)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   408
    apply (force simp add: m_assoc subsetD r_coset_eq)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   409
   apply (rule inj_on_g, assumption+)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   410
  apply (force simp add: m_assoc subsetD r_coset_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   411
 txt{*The sets @{term "H #> a"} and @{term "H"} are finite.*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   412
 apply (simp add: r_coset_subset_G [THEN finite_subset])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   413
apply (blast intro: finite_subset)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   414
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   415
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   416
subsection{*Two distinct right cosets are disjoint*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   417
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   418
lemma (in coset) rcos_equation:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   419
     "[|subgroup H G;  a \<in> carrier G;  b \<in> carrier G;  ha \<otimes> a = h \<otimes> b;
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   420
        h \<in> H;  ha \<in> H;  hb \<in> H|]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   421
      ==> \<exists>h\<in>H. h \<otimes> b = hb \<otimes> a"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   422
apply (rule bexI [of _"hb \<otimes> ((inv ha) \<otimes> h)"])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   423
apply (simp  add: m_assoc transpose_inv subgroup.subset [THEN subsetD])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   424
apply (simp add: subgroup.m_closed subgroup.m_inv_closed)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   425
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   426
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   427
lemma (in coset) rcos_disjoint:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   428
     "[|subgroup H G; a \<in> rcosets G H; b \<in> rcosets G H; a\<noteq>b|] ==> a \<inter> b = {}"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   429
apply (simp add: setrcos_eq r_coset_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   430
apply (blast intro: rcos_equation sym)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   431
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   432
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   433
lemma (in coset) setrcos_subset_PowG:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   434
     "subgroup H G  ==> rcosets G H <= Pow(carrier G)"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   435
apply (simp add: setrcos_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   436
apply (blast dest: r_coset_subset_G subgroup.subset)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   437
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   438
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   439
subsection {*Factorization of a Group*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   440
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   441
constdefs
13936
d3671b878828 Greatly extended CRing. Added Module.
ballarin
parents: 13889
diff changeset
   442
  FactGroup :: "[('a,'b) monoid_scheme, 'a set] => ('a set) monoid"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   443
     (infixl "Mod" 60)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   444
  "FactGroup G H ==
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   445
    (| carrier = rcosets G H,
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   446
       mult = (%X: rcosets G H. %Y: rcosets G H. set_mult G X Y),
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   447
       one = H (*,
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   448
       m_inv = (%X: rcosets G H. set_inv G X) *) |)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   449
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   450
lemma (in coset) setmult_closed:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   451
     "[| H <| G; K1 \<in> rcosets G H; K2 \<in> rcosets G H |]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   452
      ==> K1 <#> K2 \<in> rcosets G H"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   453
by (auto simp add: normal_imp_subgroup [THEN subgroup.subset]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   454
                   rcos_sum setrcos_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   455
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   456
lemma (in group) setinv_closed:
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   457
     "[| H <| G; K \<in> rcosets G H |] ==> set_inv G K \<in> rcosets G H"
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   458
by (auto simp add:  normal_imp_subgroup
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   459
                 subgroup.subset coset.rcos_inv [OF is_coset]
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   460
                 coset.setrcos_eq [OF is_coset])
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   461
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   462
(*
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   463
The old version is no longer valid: "group G" has to be an explicit premise.
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   464
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   465
lemma setinv_closed:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   466
     "[| H <| G; K \<in> rcosets G H |] ==> set_inv G K \<in> rcosets G H"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   467
by (auto simp add:  normal_imp_subgroup
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   468
                   subgroup.subset coset.rcos_inv coset.setrcos_eq)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   469
*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   470
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   471
lemma (in coset) setrcos_assoc:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   472
     "[|H <| G; M1 \<in> rcosets G H; M2 \<in> rcosets G H; M3 \<in> rcosets G H|]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   473
      ==> M1 <#> M2 <#> M3 = M1 <#> (M2 <#> M3)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   474
by (auto simp add: setrcos_eq rcos_sum normal_imp_subgroup
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   475
                   subgroup.subset m_assoc)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   476
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   477
lemma (in group) subgroup_in_rcosets:
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   478
  "subgroup H G ==> H \<in> rcosets G H"
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   479
proof -
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   480
  assume sub: "subgroup H G"
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   481
  have "r_coset G H \<one> = H"
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   482
    by (rule coset.coset_join2)
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   483
      (auto intro: sub subgroup.one_closed is_coset)
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   484
  then show ?thesis
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   485
    by (auto simp add: coset.setrcos_eq [OF is_coset])
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   486
qed
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   487
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   488
(*
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   489
lemma subgroup_in_rcosets:
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   490
  "subgroup H G ==> H \<in> rcosets G H"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   491
apply (frule subgroup_imp_coset)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   492
apply (frule subgroup_imp_group)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   493
apply (simp add: coset.setrcos_eq)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   494
apply (blast del: equalityI
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   495
             intro!: group.subgroup.one_closed group.one_closed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   496
                     coset.coset_join2 [symmetric])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   497
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   498
*)
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
lemma (in coset) setrcos_inv_mult_group_eq:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   501
     "[|H <| G; M \<in> rcosets G H|] ==> set_inv G M <#> M = H"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   502
by (auto simp add: setrcos_eq rcos_inv rcos_sum normal_imp_subgroup
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   503
                   subgroup.subset)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents: 13936
diff changeset
   504
(*
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   505
lemma (in group) factorgroup_is_magma:
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   506
  "H <| G ==> magma (G Mod H)"
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   507
  by rule (simp add: FactGroup_def coset.setmult_closed [OF is_coset])
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   508
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   509
lemma (in group) factorgroup_semigroup_axioms:
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   510
  "H <| G ==> semigroup_axioms (G Mod H)"
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   511
  by rule (simp add: FactGroup_def coset.setrcos_assoc [OF is_coset]
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   512
    coset.setmult_closed [OF is_coset])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents: 13936
diff changeset
   513
*)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   514
theorem (in group) factorgroup_is_group:
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   515
  "H <| G ==> group (G Mod H)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   516
apply (insert is_coset)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   517
apply (simp add: FactGroup_def)
13936
d3671b878828 Greatly extended CRing. Added Module.
ballarin
parents: 13889
diff changeset
   518
apply (rule groupI)
d3671b878828 Greatly extended CRing. Added Module.
ballarin
parents: 13889
diff changeset
   519
    apply (simp add: coset.setmult_closed)
d3671b878828 Greatly extended CRing. Added Module.
ballarin
parents: 13889
diff changeset
   520
   apply (simp add: normal_imp_subgroup subgroup_in_rcosets)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   521
  apply (simp add: restrictI coset.setmult_closed coset.setrcos_assoc)
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   522
 apply (simp add: normal_imp_subgroup
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   523
   subgroup_in_rcosets coset.setrcos_mult_eq)
13936
d3671b878828 Greatly extended CRing. Added Module.
ballarin
parents: 13889
diff changeset
   524
apply (auto dest: coset.setrcos_inv_mult_group_eq simp add: setinv_closed)
13889
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   525
done
6676ac2527fa Fixed Coset.thy (proved theorem factorgroup_is_group).
ballarin
parents: 13870
diff changeset
   526
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   527
end