src/HOL/GroupTheory/Coset.thy
author paulson
Fri, 27 Sep 2002 13:24:29 +0200
changeset 13594 c2ee8f5a5652
parent 13583 5fcc8bf538ee
child 13595 7e6cdcd113a2
permissions -rw-r--r--
Isar experiments, etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     1
(*  Title:      HOL/GroupTheory/Coset
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     2
    ID:         $Id$
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     3
    Author:     Florian Kammueller, with new proofs by L C Paulson
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     4
*)
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     5
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
     6
header{*Theory of Cosets*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
     7
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
     8
theory Coset = Group + Exponent:
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     9
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    10
constdefs
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    11
  r_coset    :: "[('a,'b) group_scheme,'a set, 'a] => 'a set"    
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    12
   "r_coset G H a == (% x. (sum G) x a) ` H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    13
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    14
  l_coset    :: "[('a,'b) group_scheme, 'a, 'a set] => 'a set"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    15
   "l_coset G a H == (% x. (sum G) a x) ` H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    16
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    17
  rcosets  :: "[('a,'b) group_scheme,'a set] => ('a set)set"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    18
   "rcosets G H == r_coset G H ` (carrier G)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    19
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    20
  set_sum  :: "[('a,'b) group_scheme,'a set,'a set] => 'a set"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    21
   "set_sum G H K == (%(x,y). (sum G) x y) ` (H \<times> K)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    22
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    23
  set_minus   :: "[('a,'b) group_scheme,'a set] => 'a set"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    24
   "set_minus G H == (gminus G) ` H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    25
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    26
  normal     :: "['a set, ('a,'b) group_scheme] => bool"       (infixl "<|" 60) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    27
   "normal H G == subgroup H G & 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    28
                  (\<forall>x \<in> carrier G. r_coset G H x = l_coset G x H)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    29
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    30
syntax (xsymbols)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    31
  normal  :: "['a set, ('a,'b) group_scheme] => bool" (infixl "\<lhd>" 60)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    32
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    33
locale coset = group G +
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    34
  fixes rcos      :: "['a set, 'a] => 'a set"     (infixl "#>" 60)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    35
    and lcos      :: "['a, 'a set] => 'a set"     (infixl "<#" 60)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    36
    and setsum   :: "['a set, 'a set] => 'a set" (infixl "<#>" 60)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    37
  defines rcos_def: "H #> x == r_coset G H x"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    38
      and lcos_def: "x <# H == l_coset G x H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    39
      and setsum_def: "H <#> K == set_sum G H K"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    40
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    41
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    42
text{*Lemmas useful for Sylow's Theorem*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    43
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    44
lemma card_inj:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    45
     "[|f \<in> A\<rightarrow>B; inj_on f A; finite A; finite B|] ==> card(A) <= card(B)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    46
apply (rule card_inj_on_le)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    47
apply (auto simp add: Pi_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    48
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    49
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    50
lemma card_bij: 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    51
     "[|f \<in> A\<rightarrow>B; inj_on f A; 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    52
        g \<in> B\<rightarrow>A; inj_on g B; finite A; finite B|] ==> card(A) = card(B)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    53
by (blast intro: card_inj order_antisym) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    54
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    55
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    56
subsection{*Lemmas Using Locale Constants*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    57
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    58
lemma (in coset) r_coset_eq: "H #> a = {b . \<exists>h\<in>H. h \<oplus> a = b}"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    59
by (auto simp add: rcos_def r_coset_def sum_def)
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    60
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    61
lemma (in coset) l_coset_eq: "a <# H = {b . \<exists>h\<in>H. a \<oplus> h = b}"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    62
by (auto simp add: lcos_def l_coset_def sum_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    63
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    64
lemma (in coset) setrcos_eq: "rcosets G H = {C . \<exists>a\<in>carrier G. C = H #> a}"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    65
by (auto simp add: rcosets_def rcos_def sum_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    66
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    67
lemma (in coset) set_sum_eq: "H <#> K = {c . \<exists>h\<in>H. \<exists>k\<in>K. c = h \<oplus> k}"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    68
by (simp add: setsum_def set_sum_def sum_def image_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    69
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    70
lemma (in coset) coset_sum_assoc:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    71
     "[| M <= carrier G; g \<in> carrier G; h \<in> carrier G |]  
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    72
      ==> (M #> g) #> h = M #> (g \<oplus> h)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    73
by (force simp add: r_coset_eq sum_assoc)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    74
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    75
lemma (in coset) coset_sum_zero [simp]: "M <= carrier G ==> M #> \<zero> = M"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    76
by (force simp add: r_coset_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    77
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    78
lemma (in coset) coset_sum_minus1:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    79
     "[| M #> (x \<oplus> (\<ominus>y)) = M;  x \<in> carrier G ; y \<in> carrier G; 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    80
         M <= carrier G |] ==> M #> x = M #> y"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    81
apply (erule subst [of concl: "%z. M #> x = z #> y"])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    82
apply (simp add: coset_sum_assoc sum_assoc)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    83
done
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    84
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    85
lemma (in coset) coset_sum_minus2:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    86
     "[| M #> x = M #> y;  x \<in> carrier G;  y \<in> carrier G;  M <= carrier G |]  
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    87
      ==> M #> (x \<oplus> (\<ominus>y)) = M "
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    88
apply (simp add: coset_sum_assoc [symmetric])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    89
apply (simp add: coset_sum_assoc)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    90
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    91
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    92
lemma (in coset) coset_join1:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    93
     "[| H #> x = H;  x \<in> carrier G;  subgroup H G |] ==> x\<in>H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    94
apply (erule subst)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    95
apply (simp add: r_coset_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    96
apply (blast intro: left_zero subgroup_zero_closed)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    97
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
    98
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
    99
text{*Locales don't currently work with @{text rule_tac}, so we
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   100
must prove this lemma separately.*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   101
lemma (in coset) solve_equation:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   102
    "\<lbrakk>subgroup H G; x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> \<exists>h\<in>H. h \<oplus> x = y"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   103
apply (rule bexI [of _ "y \<oplus> (\<ominus>x)"])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   104
apply (auto simp add: subgroup_sum_closed subgroup_minus_closed sum_assoc 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   105
                      subgroup_imp_subset [THEN subsetD])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   106
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   107
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   108
lemma (in coset) coset_join2:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   109
     "[| x \<in> carrier G;  subgroup H G;  x\<in>H |] ==> H #> x = H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   110
by (force simp add: subgroup_sum_closed r_coset_eq solve_equation)
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   111
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   112
lemma (in coset) r_coset_subset_G:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   113
     "[| H <= carrier G; x \<in> carrier G |] ==> H #> x <= carrier G"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   114
by (auto simp add: r_coset_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   115
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   116
lemma (in coset) rcosI:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   117
     "[| h \<in> H; H <= carrier G; x \<in> carrier G|] ==> h \<oplus> x \<in> H #> x"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   118
by (auto simp add: r_coset_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   119
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   120
lemma (in coset) setrcosI:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   121
     "[| H <= carrier G; x \<in> carrier G |] ==> H #> x \<in> rcosets G H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   122
by (auto simp add: setrcos_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   123
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   124
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   125
text{*Really needed?*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   126
lemma (in coset) transpose_minus:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   127
     "[| x \<oplus> y = z;  x \<in> carrier G;  y \<in> carrier G;  z \<in> carrier G |]  
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   128
      ==> (\<ominus>x) \<oplus> z = y"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   129
by (force simp add: sum_assoc [symmetric])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   130
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   131
lemma (in coset) repr_independence:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   132
     "[| y \<in> H #> x;  x \<in> carrier G; subgroup H G |] ==> H #> x = H #> y"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   133
by (auto simp add: r_coset_eq sum_assoc [symmetric] right_cancellation_iff
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   134
                   subgroup_imp_subset [THEN subsetD]
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   135
                   subgroup_sum_closed solve_equation)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   136
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   137
lemma (in coset) rcos_self: "[| x \<in> carrier G; subgroup H G |] ==> x \<in> H #> x"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   138
apply (simp add: r_coset_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   139
apply (blast intro: left_zero subgroup_imp_subset [THEN subsetD] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   140
                    subgroup_zero_closed)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   141
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   142
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   143
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   144
subsection{*normal subgroups*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   145
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   146
text{*Allows use of theorems proved in the locale coset*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   147
lemma subgroup_imp_coset: "subgroup H G ==> coset G"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   148
apply (drule subgroup_imp_group)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   149
apply (simp add: group_def coset_def)  
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   150
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   151
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   152
lemma normal_imp_subgroup: "H <| G ==> subgroup H G"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   153
by (simp add: normal_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   154
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   155
lemmas normal_imp_group = normal_imp_subgroup [THEN subgroup_imp_group]
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   156
lemmas normal_imp_coset = normal_imp_subgroup [THEN subgroup_imp_coset]
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   157
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   158
lemma (in coset) normal_iff:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   159
     "(H <| G) = (subgroup H G & (\<forall>x \<in> carrier G. H #> x = x <# H))"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   160
by (simp add: lcos_def rcos_def normal_def)
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   161
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   162
lemma (in coset) normal_imp_rcos_eq_lcos:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   163
     "[| H <| G; x \<in> carrier G |] ==> H #> x = x <# H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   164
by (simp add: lcos_def rcos_def normal_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   165
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   166
lemma (in coset) normal_minus_op_closed1:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   167
     "\<lbrakk>H \<lhd> G; x \<in> carrier G; h \<in> H\<rbrakk> \<Longrightarrow> (\<ominus>x) \<oplus> h \<oplus> x \<in> H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   168
apply (auto simp add: l_coset_eq r_coset_eq normal_iff)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   169
apply (drule bspec, assumption) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   170
apply (drule equalityD1 [THEN subsetD], blast, clarify)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   171
apply (simp add: sum_assoc subgroup_imp_subset [THEN subsetD])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   172
apply (erule subst)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   173
apply (simp add: sum_assoc [symmetric] subgroup_imp_subset [THEN subsetD])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   174
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   175
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   176
lemma (in coset) normal_minus_op_closed2:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   177
     "\<lbrakk>H \<lhd> G; x \<in> carrier G; h \<in> H\<rbrakk> \<Longrightarrow> x \<oplus> h \<oplus> (\<ominus>x) \<in> H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   178
apply (drule normal_minus_op_closed1 [of H "(\<ominus>x)"]) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   179
apply auto
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   180
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   181
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   182
lemma (in coset) lcos_sum_assoc:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   183
     "[| M <= carrier G; g \<in> carrier G; h \<in> carrier G |]  
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   184
      ==> g <# (h <# M) = (g \<oplus> h) <# M"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   185
by (force simp add: l_coset_eq sum_assoc)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   186
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   187
lemma (in coset) lcos_sum_zero: "M <= carrier G ==> \<zero> <# M = M"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   188
by (force simp add: l_coset_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   189
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   190
lemma (in coset) l_coset_subset_G:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   191
     "[| H <= carrier G; x \<in> carrier G |] ==> x <# H <= carrier G"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   192
by (auto simp add: l_coset_eq subsetD)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   193
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   194
lemma (in coset) l_repr_independence:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   195
     "[| y \<in> x <# H;  x \<in> carrier G;  subgroup H G |] ==> x <# H = y <# H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   196
apply (auto simp add: l_coset_eq sum_assoc left_cancellation_iff 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   197
                      subgroup_imp_subset [THEN subsetD] subgroup_sum_closed)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   198
apply (rule_tac x = "sum G (gminus G h) ha" in bexI)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   199
  --{*FIXME: locales don't currently work with @{text rule_tac},
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   200
      want @{term "(\<ominus>h) \<oplus> ha"}*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   201
apply (auto simp add: sum_assoc [symmetric] subgroup_imp_subset [THEN subsetD] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   202
                      subgroup_minus_closed subgroup_sum_closed)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   203
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   204
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   205
lemma (in coset) setsum_subset_G:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   206
     "[| H <= carrier G; K <= carrier G |] ==> H <#> K <= carrier G"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   207
by (auto simp add: set_sum_eq subsetD)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   208
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   209
lemma (in coset) subgroup_sum_id: "subgroup H G ==> H <#> H = H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   210
apply (auto simp add: subgroup_sum_closed set_sum_eq Sigma_def image_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   211
apply (rule_tac x = x in bexI)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   212
apply (rule bexI [of _ "\<zero>"])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   213
apply (auto simp add: subgroup_sum_closed subgroup_zero_closed 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   214
                      right_zero subgroup_imp_subset [THEN subsetD])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   215
done
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   216
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   217
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   218
(* set of inverses of an r_coset *)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   219
lemma (in coset) rcos_minus:
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   220
  assumes normalHG: "H <| G"
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   221
      and xinG:     "x \<in> carrier G"
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   222
  shows "set_minus G (H #> x) = H #> (\<ominus>x)"
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   223
proof -
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   224
  have H_subset: "H <= carrier G" 
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   225
    by (rule subgroup_imp_subset [OF normal_imp_subgroup, OF normalHG])
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   226
  show ?thesis
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   227
  proof (auto simp add: r_coset_eq image_def set_minus_def)
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   228
    fix h
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   229
    assume "h \<in> H"
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   230
      hence "((\<ominus>x) \<oplus> (\<ominus>h) \<oplus> x) \<oplus> \<ominus>x = \<ominus>(h \<oplus> x)"
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   231
	by (simp add: xinG sum_assoc minus_sum H_subset [THEN subsetD])
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   232
      thus "\<exists>ha\<in>H. ha \<oplus> \<ominus>x = \<ominus>(h \<oplus> x)" 
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   233
       using prems
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   234
	by (blast intro: normal_minus_op_closed1 normal_imp_subgroup 
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   235
			 subgroup_minus_closed)
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   236
  next
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   237
    fix h
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   238
    assume "h \<in> H"
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   239
      hence eq: "(x \<oplus> (\<ominus>h) \<oplus> (\<ominus>x)) \<oplus> x = x \<oplus> \<ominus>h"
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   240
        by (simp add: xinG sum_assoc H_subset [THEN subsetD])
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   241
      hence "(\<exists>j\<in>H. j \<oplus> x = \<ominus> (h \<oplus> (\<ominus>x))) \<and> h \<oplus> \<ominus>x = \<ominus>(\<ominus>(h \<oplus> (\<ominus>x)))"
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   242
       using prems
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   243
	by (simp add: sum_assoc minus_sum H_subset [THEN subsetD],
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   244
            blast intro: eq normal_minus_op_closed2 normal_imp_subgroup 
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   245
			 subgroup_minus_closed)
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   246
      thus "\<exists>y. (\<exists>h\<in>H. h \<oplus> x = y) \<and> h \<oplus> \<ominus>x = \<ominus>y" ..
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   247
  qed
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   248
qed
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   249
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   250
(*The old proof is something like this, but the rule_tac calls make
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   251
illegal references to implicit structures.
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   252
lemma (in coset) rcos_minus:
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   253
     "[| H <| G; x \<in> carrier G |] ==> set_minus G (H #> x) = H #> (\<ominus>x)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   254
apply (frule normal_imp_subgroup)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   255
apply (auto simp add: r_coset_eq image_def set_minus_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   256
apply (rule_tac x = "(\<ominus>x) \<oplus> (\<ominus>h) \<oplus> x" in bexI)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   257
apply (simp add: sum_assoc minus_sum subgroup_imp_subset [THEN subsetD])
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   258
apply (simp add: subgroup_minus_closed, assumption+)
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   259
apply (rule_tac x = "\<ominus> (h \<oplus> (\<ominus>x)) " in exI)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   260
apply (simp add: minus_sum subgroup_imp_subset [THEN subsetD])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   261
apply (rule_tac x = "x \<oplus> (\<ominus>h) \<oplus> (\<ominus>x)" in bexI)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   262
apply (simp add: sum_assoc subgroup_imp_subset [THEN subsetD])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   263
apply (simp add: normal_minus_op_closed2 subgroup_minus_closed)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   264
done
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   265
*)
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   266
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   267
lemma (in coset) rcos_minus2:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   268
     "[| H <| G; K \<in> rcosets G H; x \<in> K |] ==> set_minus G K = H #> (\<ominus>x)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   269
apply (simp add: setrcos_eq, clarify)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   270
apply (subgoal_tac "x : carrier G")
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   271
 prefer 2
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   272
 apply (blast dest: r_coset_subset_G subgroup_imp_subset normal_imp_subgroup) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   273
apply (drule repr_independence)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   274
  apply assumption
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   275
 apply (erule normal_imp_subgroup)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   276
apply (simp add: rcos_minus)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   277
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   278
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   279
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   280
(* some rules for <#> with #> or <# *)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   281
lemma (in coset) setsum_rcos_assoc:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   282
     "[| H <= carrier G; K <= carrier G; x \<in> carrier G |] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   283
      ==> H <#> (K #> x) = (H <#> K) #> x"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   284
apply (auto simp add: rcos_def r_coset_def setsum_def set_sum_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   285
apply (force simp add: sum_assoc)+
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   286
done
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   287
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   288
lemma (in coset) rcos_assoc_lcos:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   289
     "[| H <= carrier G; K <= carrier G; x \<in> carrier G |] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   290
      ==> (H #> x) <#> K = H <#> (x <# K)"
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   291
apply (auto simp add: rcos_def r_coset_def lcos_def l_coset_def 
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   292
                      setsum_def set_sum_def Sigma_def image_def)
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   293
apply (force intro!: exI bexI simp add: sum_assoc)+
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   294
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   295
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   296
lemma (in coset) rcos_sum_step1:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   297
     "[| H <| G; x \<in> carrier G; y \<in> carrier G |] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   298
      ==> (H #> x) <#> (H #> y) = (H <#> (x <# H)) #> y"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   299
by (simp add: setsum_rcos_assoc normal_imp_subgroup [THEN subgroup_imp_subset]
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   300
              r_coset_subset_G l_coset_subset_G rcos_assoc_lcos)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   301
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   302
lemma (in coset) rcos_sum_step2:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   303
     "[| H <| G; x \<in> carrier G; y \<in> carrier G |] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   304
      ==> (H <#> (x <# H)) #> y = (H <#> (H #> x)) #> y"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   305
by (simp add: normal_imp_rcos_eq_lcos)
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   306
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   307
lemma (in coset) rcos_sum_step3:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   308
     "[| H <| G; x \<in> carrier G; y \<in> carrier G |] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   309
      ==> (H <#> (H #> x)) #> y = H #> (x \<oplus> y)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   310
by (simp add: setsum_rcos_assoc r_coset_subset_G coset_sum_assoc
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   311
              setsum_subset_G  subgroup_sum_id
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   312
              subgroup_imp_subset normal_imp_subgroup)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   313
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   314
lemma (in coset) rcos_sum:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   315
     "[| H <| G; x \<in> carrier G; y \<in> carrier G |] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   316
      ==> (H #> x) <#> (H #> y) = H #> (x \<oplus> y)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   317
by (simp add: rcos_sum_step1 rcos_sum_step2 rcos_sum_step3)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   318
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   319
(*generalizes subgroup_sum_id*)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   320
lemma (in coset) setrcos_sum_eq: "[|H <| G; M \<in> rcosets G H|] ==> H <#> M = M"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   321
by (auto simp add: setrcos_eq normal_imp_subgroup subgroup_imp_subset
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   322
                   setsum_rcos_assoc subgroup_sum_id)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   323
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   324
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   325
subsection{*Lemmas Leading to Lagrange's Theorem*}
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   326
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   327
lemma (in coset) setrcos_part_G: "subgroup H G ==> \<Union> rcosets G H = carrier G"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   328
apply (rule equalityI)
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   329
apply (force simp add: subgroup_imp_subset [THEN subsetD] 
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   330
                       setrcos_eq r_coset_eq)
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   331
apply (auto simp add: setrcos_eq subgroup_imp_subset rcos_self)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   332
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   333
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   334
lemma (in coset) cosets_finite:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   335
     "[| c \<in> rcosets G H;  H <= carrier G;  finite (carrier G) |] ==> finite c"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   336
apply (auto simp add: setrcos_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   337
apply (simp (no_asm_simp) add: r_coset_subset_G [THEN finite_subset])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   338
done
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   339
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   340
lemma (in coset) card_cosets_equal:
13594
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   341
     "[| c \<in> rcosets G H;  H <= carrier G; finite(carrier G) |] 
c2ee8f5a5652 Isar experiments, etc.
paulson
parents: 13583
diff changeset
   342
      ==> card c = card H"
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   343
apply (auto simp add: setrcos_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   344
(*FIXME: I can't say
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   345
apply (rule_tac f = "%y. y \<oplus> (\<ominus>a)" and g = "%y. y \<oplus> a" in card_bij_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   346
*)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   347
apply (rule_tac f = "%y. sum G y (gminus G a)" 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   348
            and g = "%y. sum G y a" in card_bij_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   349
     apply (simp add: r_coset_subset_G [THEN finite_subset])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   350
    apply (blast intro: finite_subset)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   351
   txt{*inclusion in @{term H}*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   352
   apply (force simp add: sum_assoc subsetD r_coset_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   353
  defer 1
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   354
   txt{*inclusion in @{term "H #> a"}*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   355
  apply (blast intro: rcosI)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   356
 apply (force simp add: inj_on_def right_cancellation_iff subsetD)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   357
apply (rule inj_onI)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   358
apply (subgoal_tac "x \<in> carrier G & y \<in> carrier G")
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   359
 prefer 2 apply (blast intro: r_coset_subset_G [THEN subsetD])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   360
apply (rotate_tac -1)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   361
apply (simp add: right_cancellation_iff subsetD)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   362
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   363
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   364
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   365
subsection{*Two distinct right cosets are disjoint*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   366
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   367
lemma (in coset) rcos_equation:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   368
     "[|subgroup H G;  a \<in> carrier G;  b \<in> carrier G;  ha \<oplus> a = h \<oplus> b;   
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   369
        h \<in> H;  ha \<in> H;  hb \<in> H|]  
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   370
      ==> \<exists>h\<in>H. h \<oplus> b = hb \<oplus> a"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   371
apply (rule bexI [of _"hb \<oplus> ((\<ominus>ha) \<oplus> h)"])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   372
apply (simp  add: sum_assoc transpose_minus subgroup_imp_subset [THEN subsetD])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   373
apply (simp add: subgroup_sum_closed subgroup_minus_closed)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   374
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   375
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   376
lemma (in coset) rcos_disjoint:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   377
     "[|subgroup H G; a \<in> rcosets G H; b \<in> rcosets G H; a\<noteq>b|] ==> a \<inter> b = {}"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   378
apply (simp add: setrcos_eq r_coset_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   379
apply (blast intro: rcos_equation sym)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   380
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   381
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   382
lemma (in coset) setrcos_subset_PowG:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   383
     "subgroup H G  ==> rcosets G H <= Pow(carrier G)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   384
apply (simp add: setrcos_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   385
apply (blast dest: r_coset_subset_G subgroup_imp_subset)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   386
done
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   387
13583
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   388
theorem (in coset) lagrange:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   389
     "[| finite(carrier G); subgroup H G |] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   390
      ==> card(rcosets G H) * card(H) = order(G)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   391
apply (simp (no_asm_simp) add: order_def setrcos_part_G [symmetric])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   392
apply (subst mult_commute)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   393
apply (rule card_partition)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   394
   apply (simp add: setrcos_subset_PowG [THEN finite_subset])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   395
  apply (simp add: setrcos_part_G)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   396
 apply (simp add: card_cosets_equal subgroup_def)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   397
apply (simp add: rcos_disjoint)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   398
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   399
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   400
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   401
subsection {*Factorization of a Group*}
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   402
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   403
constdefs
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   404
  FactGroup :: "[('a,'b) group_scheme, 'a set] => ('a set) group"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   405
     (infixl "Mod" 60)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   406
   "FactGroup G H ==
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   407
      (| carrier = rcosets G H,
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   408
	 sum = (%X: rcosets G H. %Y: rcosets G H. set_sum G X Y),
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   409
	 gminus = (%X: rcosets G H. set_minus G X), 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   410
	 zero = H |)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   411
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   412
lemma (in coset) setsum_closed:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   413
     "[| H <| G; K1 \<in> rcosets G H; K2 \<in> rcosets G H |] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   414
      ==> K1 <#> K2 \<in> rcosets G H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   415
by (auto simp add: normal_imp_subgroup [THEN subgroup_imp_subset] 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   416
                   rcos_sum setrcos_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   417
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   418
lemma setminus_closed:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   419
     "[| H <| G; K \<in> rcosets G H |] ==> set_minus G K \<in> rcosets G H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   420
by (auto simp add: normal_imp_coset normal_imp_group normal_imp_subgroup
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   421
                   subgroup_imp_subset coset.rcos_minus coset.setrcos_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   422
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   423
lemma (in coset) setrcos_assoc:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   424
     "[|H <| G; M1 \<in> rcosets G H; M2 \<in> rcosets G H; M3 \<in> rcosets G H|]   
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   425
      ==> M1 <#> M2 <#> M3 = M1 <#> (M2 <#> M3)"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   426
by (auto simp add: setrcos_eq rcos_sum normal_imp_subgroup 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   427
                   subgroup_imp_subset sum_assoc)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   428
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   429
lemma subgroup_in_rcosets: "subgroup H G ==> H \<in> rcosets G H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   430
apply (frule subgroup_imp_coset) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   431
apply (frule subgroup_imp_group) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   432
apply (simp add: coset.setrcos_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   433
apply (blast del: equalityI 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   434
             intro!: group.subgroup_zero_closed group.zero_closed
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   435
                     coset.coset_join2 [symmetric])
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   436
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   437
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   438
lemma (in coset) setrcos_minus_sum_eq:
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   439
     "[|H <| G; M \<in> rcosets G H|] ==> set_minus G M <#> M = H"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   440
by (auto simp add: setrcos_eq rcos_minus rcos_sum normal_imp_subgroup 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   441
                   subgroup_imp_subset)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   442
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   443
theorem factorgroup_is_group: "H <| G ==> (G Mod H) \<in> Group"
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   444
apply (frule normal_imp_coset) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   445
apply (simp add: FactGroup_def) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   446
apply (rule group.intro)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   447
 apply (rule semigroup.intro) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   448
  apply (simp add: restrictI coset.setsum_closed) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   449
 apply (simp add: coset.setsum_closed coset.setrcos_assoc)  
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   450
apply (rule group_axioms.intro) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   451
   apply (simp add: restrictI setminus_closed) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   452
  apply (simp add: normal_imp_subgroup subgroup_in_rcosets) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   453
 apply (simp add: setminus_closed coset.setrcos_minus_sum_eq) 
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   454
apply (simp add: normal_imp_subgroup subgroup_in_rcosets coset.setrcos_sum_eq)
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   455
done
5fcc8bf538ee converted to Isar and using new "implicit structures" instead of Sigma, etc
paulson
parents: 11394
diff changeset
   456
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   457
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
   458
end