src/HOL/Algebra/Sylow.thy
changeset 14963 d584e32f7d46
parent 14803 f7557773cc87
child 16417 9bc16273c2d4
     1.1 --- a/src/HOL/Algebra/Sylow.thy	Thu Jun 17 14:27:01 2004 +0200
     1.2 +++ b/src/HOL/Algebra/Sylow.thy	Thu Jun 17 17:18:30 2004 +0200
     1.3 @@ -206,7 +206,7 @@
     1.4  
     1.5  lemma (in sylow_central) M1_cardeq_rcosetGM1g:
     1.6       "g \<in> carrier G ==> card(M1 #> g) = card(M1)"
     1.7 -by (simp (no_asm_simp) add: M1_subset_G card_cosets_equal setrcosI)
     1.8 +by (simp (no_asm_simp) add: M1_subset_G card_cosets_equal rcosetsI)
     1.9  
    1.10  lemma (in sylow_central) M1_RelM_rcosetGM1g:
    1.11       "g \<in> carrier G ==> (M1, M1 #> g) \<in> RelM"
    1.12 @@ -219,16 +219,14 @@
    1.13  done
    1.14  
    1.15  
    1.16 +subsection{*Equal Cardinalities of @{term M} and the Set of Cosets*}
    1.17  
    1.18 -subsection{*Equal Cardinalities of @{term M} and @{term "rcosets G H"}*}
    1.19 -
    1.20 -text{*Injections between @{term M} and @{term "rcosets G H"} show that
    1.21 +text{*Injections between @{term M} and @{term "rcosets\<^bsub>G\<^esub> H"} show that
    1.22   their cardinalities are equal.*}
    1.23  
    1.24  lemma ElemClassEquiv:
    1.25 -     "[| equiv A r; C\<in>A // r |] ==> \<forall>x \<in> C. \<forall>y \<in> C. (x,y)\<in>r"
    1.26 -apply (unfold equiv_def quotient_def sym_def trans_def, blast)
    1.27 -done
    1.28 +     "[| equiv A r; C \<in> A // r |] ==> \<forall>x \<in> C. \<forall>y \<in> C. (x,y)\<in>r"
    1.29 +by (unfold equiv_def quotient_def sym_def trans_def, blast)
    1.30  
    1.31  lemma (in sylow_central) M_elem_map:
    1.32       "M2 \<in> M ==> \<exists>g. g \<in> carrier G & M1 #> g = M2"
    1.33 @@ -243,16 +241,16 @@
    1.34  lemmas (in sylow_central) M_elem_map_eq =
    1.35          M_elem_map [THEN someI_ex, THEN conjunct2]
    1.36  
    1.37 -lemma (in sylow_central) M_funcset_setrcos_H:
    1.38 -     "(%x:M. H #> (SOME g. g \<in> carrier G & M1 #> g = x)) \<in> M \<rightarrow> rcosets G H"
    1.39 -apply (rule setrcosI [THEN restrictI])
    1.40 +lemma (in sylow_central) M_funcset_rcosets_H:
    1.41 +     "(%x:M. H #> (SOME g. g \<in> carrier G & M1 #> g = x)) \<in> M \<rightarrow> rcosets H"
    1.42 +apply (rule rcosetsI [THEN restrictI])
    1.43  apply (rule H_is_subgroup [THEN subgroup.subset])
    1.44  apply (erule M_elem_map_carrier)
    1.45  done
    1.46  
    1.47 -lemma (in sylow_central) inj_M_GmodH: "\<exists>f \<in> M\<rightarrow>rcosets G H. inj_on f M"
    1.48 +lemma (in sylow_central) inj_M_GmodH: "\<exists>f \<in> M\<rightarrow>rcosets H. inj_on f M"
    1.49  apply (rule bexI)
    1.50 -apply (rule_tac [2] M_funcset_setrcos_H)
    1.51 +apply (rule_tac [2] M_funcset_rcosets_H)
    1.52  apply (rule inj_onI, simp)
    1.53  apply (rule trans [OF _ M_elem_map_eq])
    1.54  prefer 2 apply assumption
    1.55 @@ -267,11 +265,11 @@
    1.56  done
    1.57  
    1.58  
    1.59 -(** the opposite injection **)
    1.60 +subsubsection{*The opposite injection*}
    1.61  
    1.62  lemma (in sylow_central) H_elem_map:
    1.63 -     "H1\<in>rcosets G H ==> \<exists>g. g \<in> carrier G & H #> g = H1"
    1.64 -by (auto simp add: setrcos_eq)
    1.65 +     "H1 \<in> rcosets H ==> \<exists>g. g \<in> carrier G & H #> g = H1"
    1.66 +by (auto simp add: RCOSETS_def)
    1.67  
    1.68  lemmas (in sylow_central) H_elem_map_carrier =
    1.69          H_elem_map [THEN someI_ex, THEN conjunct1]
    1.70 @@ -281,14 +279,13 @@
    1.71  
    1.72  
    1.73  lemma EquivElemClass:
    1.74 -     "[|equiv A r; M\<in>A // r; M1\<in>M; (M1, M2)\<in>r |] ==> M2\<in>M"
    1.75 -apply (unfold equiv_def quotient_def sym_def trans_def, blast)
    1.76 -done
    1.77 +     "[|equiv A r; M \<in> A//r; M1\<in>M; (M1,M2) \<in> r |] ==> M2 \<in> M"
    1.78 +by (unfold equiv_def quotient_def sym_def trans_def, blast)
    1.79  
    1.80 -lemma (in sylow_central) setrcos_H_funcset_M:
    1.81 -     "(\<lambda>C \<in> rcosets G H. M1 #> (@g. g \<in> carrier G \<and> H #> g = C))
    1.82 -      \<in> rcosets G H \<rightarrow> M"
    1.83 -apply (simp add: setrcos_eq)
    1.84 +
    1.85 +lemma (in sylow_central) rcosets_H_funcset_M:
    1.86 +  "(\<lambda>C \<in> rcosets H. M1 #> (@g. g \<in> carrier G \<and> H #> g = C)) \<in> rcosets H \<rightarrow> M"
    1.87 +apply (simp add: RCOSETS_def)
    1.88  apply (fast intro: someI2
    1.89              intro!: restrictI M1_in_M
    1.90                EquivElemClass [OF RelM_equiv M_in_quot _  M1_RelM_rcosetGM1g])
    1.91 @@ -296,9 +293,9 @@
    1.92  
    1.93  text{*close to a duplicate of @{text inj_M_GmodH}*}
    1.94  lemma (in sylow_central) inj_GmodH_M:
    1.95 -     "\<exists>g \<in> rcosets G H\<rightarrow>M. inj_on g (rcosets G H)"
    1.96 +     "\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)"
    1.97  apply (rule bexI)
    1.98 -apply (rule_tac [2] setrcos_H_funcset_M)
    1.99 +apply (rule_tac [2] rcosets_H_funcset_M)
   1.100  apply (rule inj_onI)
   1.101  apply (simp)
   1.102  apply (rule trans [OF _ H_elem_map_eq])
   1.103 @@ -323,10 +320,10 @@
   1.104  apply (rule calM_subset_PowG, blast)
   1.105  done
   1.106  
   1.107 -lemma (in sylow_central) cardMeqIndexH: "card(M) = card(rcosets G H)"
   1.108 +lemma (in sylow_central) cardMeqIndexH: "card(M) = card(rcosets H)"
   1.109  apply (insert inj_M_GmodH inj_GmodH_M)
   1.110  apply (blast intro: card_bij finite_M H_is_subgroup
   1.111 -             setrcos_subset_PowG [THEN finite_subset]
   1.112 +             rcosets_subset_PowG [THEN finite_subset]
   1.113               finite_Pow_iff [THEN iffD2])
   1.114  done
   1.115  
   1.116 @@ -373,3 +370,4 @@
   1.117  done
   1.118  
   1.119  end
   1.120 +