Theory Algebra3

Up to index of Isabelle/HOL/Valuation

theory Algebra3
imports Algebra2
begin

(**
                                  author Hidetsune Kobayashi
                                  Department of Mathematics
                                  Nihon University
                                  hikoba@math.cst.nihon-u.ac.jp
                                  May 3, 2004.
                                  April 6, 2007 (revised)

 chapter 3.  Group Theory. Focused on Jordan Hoelder theorem (continued) 
   section 5.    setproducts
   section 6.    preliminary lemmas for Zassenhaus
   section 7.    homomorphism
   section 8.    gkernel, kernel of a group homomorphism
   section 9.    image, image of a group homomorphism
   section 10.   incuded homomorphisms
   section 11.   Homomorphism therems
   section 12.   isomorphisms
   section 13.   Zassenhaus
   section 14.   an automorphism group
   section 15.   chain of groups I
   section 16.   existence of reduced chain 
   section 17.   existence of reduced chain and composition series"
   section 18.   chain of groups II
   section 19.   Jordan Hoelder theorem
   **)

theory Algebra3 imports Algebra2 begin


section "5. Setproducts"
       
constdefs (structure G)
  commutators:: "_ => 'a set"
     "commutators G == {z. ∃ a∈ (carrier G). ∃b ∈ (carrier G). 
                        ((a · b) · (ρ a)) · (ρ b) = z}" 

lemma (in Group) contain_commutator:"[|G » H; (commutators G) ⊆ H|] ==> G \<triangleright> H" 
apply (rule cond_nsg[of "H"], assumption)
apply (rule ballI)+
 apply (frule_tac h = h in sg_subset_elem[of "H"], assumption,
        frule_tac a = h in i_closed,
        frule_tac a = a in i_closed,
        frule_tac a = a and b = h in mult_closed, assumption+,
        frule_tac a = "a · h" and b = "ρ a" in mult_closed, assumption+)
 apply (frule_tac a = "a · h · (ρ a)" and b = "ρ h" and c = h in tassoc,
          assumption+)
 apply (simp add:l_i r_unit)
 apply (rule_tac a = "a · h · ρ a · ρ h · h" and A = H and b = "a · h · ρ a" 
        in eq_elem_in,
        rule_tac x = "a · h · ρ a · ρ h" and y = h in sg_mult_closed[of "H"],
         assumption,
        rule_tac c = "a · h · ρ a · ρ h" in subsetD[of "commutators G" "H"],
         assumption,
        thin_tac "commutators G ⊆ H",
        simp add:commutators_def, blast)
 apply assumption+
done

constdefs (structure G)
 s_top :: "[_ , 'a set, 'a set] => 'a set "  
  "s_top G H K == {z. ∃x ∈ H. ∃y ∈ K. (x · y = z)}"

syntax 
  "@S_TOP":: "['a set, ('a, 'm) Group_scheme, 'a set] => 'a set"
            ("(3_ \<struct>_ _)" [66,66,67]66)   
translations
   "H \<struct>G K" == "s_top G H K" 

lemma (in Group) s_top_induced:"[|G » L; H ⊆ L; K ⊆ L|] ==> 
                                        H \<struct>(Gp G L) K =  H \<struct>G K"
by (simp add:s_top_def Gp_def) 

lemma (in Group) s_top_l_unit:"G » K ==> {\<one>} \<struct>G K = K"
apply (rule equalityI)
 apply (rule subsetI, simp add:s_top_def, erule bexE,
        frule_tac h = y in sg_subset_elem[of "K"], assumption+,
        simp add:l_unit)
 apply (rule subsetI,
        simp add:s_top_def)
 apply (frule_tac h = x in sg_subset_elem, assumption,
        frule_tac a = x in l_unit, blast)
done

lemma (in Group) s_top_r_unit:"G » K ==> K \<struct>G {\<one>} = K" 
apply (rule equalityI)
apply (rule subsetI, simp add:s_top_def, erule bexE,
       frule_tac h = xa in sg_subset_elem[of "K"], assumption+,
       simp add:r_unit)
apply (rule subsetI,
       simp add:s_top_def,
       frule_tac h = x in sg_subset_elem[of "K"], assumption+,
       frule_tac a = x in r_unit, blast)
done

lemma (in Group) s_top_sub:"[|G » H; G » K|] ==>  H \<struct>G K ⊆ carrier G"
apply (rule subsetI) apply (simp add:s_top_def)
apply (erule bexE)+
apply (frule_tac h = xa in sg_subset_elem [of "H"], assumption+,
       frule_tac h = y in sg_subset_elem[of "K"], assumption+,
       frule_tac a = xa and b = y in mult_closed, assumption+, simp)
done

lemma (in Group) sg_inc_set_mult:"[|G » L; H ⊆ L; K ⊆ L|] ==> H \<struct>G K ⊆ L"
apply (rule subsetI)
apply (simp add:s_top_def, (erule bexE)+)
apply (frule_tac c = xa in subsetD [of "H" "L"], assumption+,
       frule_tac c = y in subsetD [of "K" "L"], assumption+,
       frule_tac x = xa and y = y in sg_mult_closed[of "L"], assumption+)
apply simp
done

lemma (in Group) s_top_sub1:"[|H ⊆ (carrier G); K ⊆ (carrier G)|] ==>  
                               H \<struct>G K ⊆ carrier G"
apply (rule subsetI)
apply (simp add:s_top_def)
apply (erule bexE)+
apply (frule_tac c = xa in subsetD[of "H" "carrier G"], assumption+,
       frule_tac c = y in subsetD[of "K" "carrier G"], assumption+,
       frule_tac a = xa and b = y in mult_closed, assumption+, simp)
done

lemma (in Group) s_top_elem:"[|G » H; G » K; a ∈ H; b ∈ K|] ==> a · b ∈ H \<struct>G K"
by (simp add:s_top_def, blast)

lemma (in Group) s_top_elem1:"[|H ⊆ carrier G; K ⊆ carrier G; a ∈ H; b ∈ K|] ==>
                    a · b ∈ H \<struct>G K "
by (simp add:s_top_def, blast)

lemma (in Group) mem_s_top:"[|H ⊆ carrier G; K ⊆ carrier G; u ∈ H \<struct>G K|] ==>
                 ∃a ∈ H. ∃b ∈ K. (a · b = u)"
by (simp add:s_top_def)

lemma (in Group) s_top_mono:"[|H ⊆ carrier G; K ⊆ carrier G; H1 ⊆ H; K1 ⊆ K|]
       ==>  H1 \<struct>G K1 ⊆ H \<struct>G K"
by (rule subsetI, simp add:s_top_def, blast)

lemma (in Group) s_top_unit_closed:"[|G » H; G » K|] ==>  \<one> ∈ H \<struct>G K"
apply (frule sg_unit_closed [of "H"], 
       frule sg_unit_closed [of "K"])
apply (cut_tac unit_closed,
       frule l_unit[of "\<one>"])
apply (simp add:s_top_def, blast)
done

lemma (in Group) s_top_commute:"[|G » H; G » K; K \<struct>G H = H \<struct>G K;
       u ∈ H \<struct>G K;  v ∈  H \<struct>G K|] ==>  u · v ∈ H \<struct>G K"
apply (frule sg_subset[of "H"], frule sg_subset[of "K"],
       frule mem_s_top[of "H" "K" "u"], assumption+, (erule bexE)+,
       frule mem_s_top[of "H" "K" "v"], assumption+, (erule bexE)+)
apply (rotate_tac 4, frule sym, thin_tac "a · b = u", frule sym, 
       thin_tac "aa · ba = v", simp,
       thin_tac "u = a · b", thin_tac "v = aa · ba")
apply (frule_tac h = a in sg_subset_elem[of "H"], assumption+,
       frule_tac h = aa in sg_subset_elem[of "H"], assumption+,
       frule_tac h = b in sg_subset_elem[of "K"], assumption+,
       frule_tac h = ba in sg_subset_elem[of "K"], assumption+)
apply (simp add:tOp_assocTr41[THEN sym], simp add:tOp_assocTr42)
apply (frule_tac a = b and b = aa in s_top_elem1[of "K" "H"], assumption+,
       simp, thin_tac "K \<struct>G H = H \<struct>G K")
apply (frule_tac u = "b · aa" in mem_s_top[of "H" "K"], assumption+,
       (erule bexE)+, frule sym, thin_tac "ab · bb = b · aa", simp,
        thin_tac "b · aa = ab · bb")
apply (frule_tac h = ab in sg_subset_elem[of "H"], assumption+,
       frule_tac h = bb in sg_subset_elem[of "K"], assumption+)
apply (simp add:tOp_assocTr42[THEN sym], simp add:tOp_assocTr41)
apply (frule_tac x = a and y = ab in sg_mult_closed[of "H"], assumption+,
       frule_tac x = bb and y = ba in sg_mult_closed[of "K"], assumption+,
       simp add:s_top_elem1)
done

lemma (in Group) s_top_commute1:"[|G » H; G » K; K \<struct>G H = H \<struct>G K;
                                        u ∈ H \<struct>G K|] ==> (ρ u) ∈ H \<struct>G K"
apply (frule sg_subset[of "H"], frule sg_subset[of "K"],
       frule mem_s_top[of "H" "K" "u"], assumption+, (erule bexE)+)
 apply (frule_tac h = a in sg_subset_elem[of "H"], assumption+,
        frule_tac h = b in sg_subset_elem[of "K"], assumption+,
        frule_tac a = a and b = b in i_ab, assumption+,
        rotate_tac 4, frule sym, thin_tac "a · b = u", simp,
        thin_tac "ρ (a · b) = ρ b · ρ a")
 apply (frule_tac x = a in sg_i_closed[of "H"], assumption+,
        frule_tac x = b in sg_i_closed[of "K"], assumption+,
        frule_tac a = "ρ b" and b = "ρ a" in s_top_elem1[of "K" "H"],
         assumption+, simp)
done

lemma (in Group) s_top_commute_sg:"[|G » H; G » K; K \<struct>G H = H \<struct>G K|] ==>
                                     G » (H \<struct>G K)"
apply (subst sg_def)
apply (frule s_top_unit_closed[of "H" "K"], assumption,
       simp add:nonempty, simp add:s_top_sub)
apply ((rule ballI)+, 
       frule_tac u = b in s_top_commute1[of "H" "K"], assumption+,
       rule_tac u = a and v = "ρ b" in s_top_commute[of "H" "K"], 
              assumption+)
done

lemma (in Group) s_top_assoc:"[|G » H; G » K; G » L|] ==>
                                 (H \<struct>G K) \<struct>G L =  H \<struct>G (K \<struct>G L)"
apply (rule equalityI)
 apply (rule subsetI, simp add:s_top_def) apply (erule exE)
 apply (erule conjE)
 apply (erule bexE)+
 apply (rotate_tac -1, frule sym, thin_tac "xb · ya = xa", simp,
        thin_tac "xa = xb · ya", frule sym, thin_tac "xb · ya · y = x",
        simp) 
apply (frule_tac h = xb in sg_subset_elem[of "H"], assumption+,
       frule_tac h = y in sg_subset_elem[of "L"], assumption+,
       frule_tac h = ya in sg_subset_elem[of "K"], assumption+,
       simp add:tassoc, blast)
apply (rule subsetI, simp add:s_top_def,
       erule bexE, erule exE, erule conjE, (erule bexE)+,
       rotate_tac -1, frule sym, thin_tac "xb · ya = y", simp,
       thin_tac "y = xb · ya")
apply (frule_tac h = xa in sg_subset_elem[of "H"], assumption+,
       frule_tac h = ya in sg_subset_elem[of "L"], assumption+,
       frule_tac h = xb in sg_subset_elem[of "K"], assumption+,
       simp add:tassoc[THEN sym],
       frule sym, thin_tac "xa · xb · ya = x", simp, blast)
done

lemma (in Group) s_topTr6:"[|G » H1; G » H2; G » K; H1 ⊆ K|] ==>
                               (H1 \<struct>G H2) ∩ K = H1 \<struct>G (H2 ∩ K)" 
apply (rule equalityI)
 apply (rule subsetI,
        simp add:s_top_def, erule conjE, (erule bexE)+,
        frule sym, thin_tac "xa · y = x", simp,
        frule_tac c = xa in subsetD[of "H1" "K"], assumption+,
        frule_tac x = "xa · y" in inEx[of _ "K"], erule bexE,
        frule_tac x = xa in sg_i_closed[of "K"], assumption+,
        frule_tac x = "ρ xa" and y = ya in sg_mult_closed[of "K"], 
        assumption+, simp)
 apply (frule_tac h = xa in sg_subset_elem[of "K"], assumption+,
        frule_tac h = "ρ xa" in sg_subset_elem[of "K"], assumption+,
        frule_tac h = y in sg_subset_elem[of "H2"], assumption+,
        simp add:tassoc[THEN sym] l_i l_unit)
 apply blast        

 apply (rule subsetI,
        simp add:s_top_def, (erule bexE)+, simp,
        frule sym, thin_tac "xa · y = x", simp,
        frule_tac c = xa in subsetD[of "H1" "K"], assumption+,
        frule_tac x = xa and y = y in sg_mult_closed[of "K"], assumption+,
        simp)
 apply blast
done

lemma (in Group) s_topTr6_1:"[|G » H1; G » H2; G » K; H2 ⊆ K|] ==>
                              (H1 \<struct>G H2) ∩ K = (H1 ∩ K) \<struct>G H2" 
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:s_top_def, erule conjE, (erule bexE)+)
apply (frule_tac  c = y in subsetD [of "H2" "K"], assumption+)
apply (frule_tac x = y in sg_i_closed [of "K"], assumption)
apply (frule_tac h = xa in sg_subset_elem[of "H1"], assumption+,
       frule_tac h = x in sg_subset_elem[of "K"], assumption+,
       frule_tac h = y in sg_subset_elem[of "K"], assumption+,
       frule_tac h = "ρ y" in sg_subset_elem[of "K"], assumption+,
       frule sym, thin_tac "xa · y = x", 
       frule_tac x = x and y = "ρ y" in sg_mult_closed[of "K"], assumption+,
       simp add:tassoc r_i r_unit, blast)

apply (rule subsetI, simp add:s_top_def, (erule bexE)+,
       simp, erule conjE,
       frule_tac c = y in subsetD[of "H2" "K"], assumption+,
       frule_tac x = xa and y = y in sg_mult_closed[of "K"], assumption+, simp,
       blast)
done

lemma (in Group) l_sub_smult:"[|G » H; G » K|] ==> H ⊆ H \<struct>G K"
apply (rule subsetI,
       simp add:s_top_def)
apply (frule sg_unit_closed[of "K"],
       frule_tac h = x in sg_subset_elem[of "H"], assumption+,
       frule_tac a = x in r_unit) 
apply blast
done

lemma (in Group) r_sub_smult:"[|G » H; G » K|] ==> K ⊆ H \<struct>G K"
apply (rule subsetI,
       simp add:s_top_def)
apply (frule sg_unit_closed[of "H"],
       frule_tac h = x in sg_subset_elem[of "K"], assumption+,
       frule_tac a = x in l_unit)
apply blast
done

lemma (in Group) s_topTr8:"G » H ==> H = H \<struct>G H"
apply (frule l_sub_smult[of "H" "H"], assumption)
apply (rule equalityI, assumption)
apply (rule subsetI)
apply (thin_tac "H ⊆ H \<struct>G H",
       simp add:s_top_def, (erule bexE)+) 
apply (frule_tac x = xa and y = y in sg_mult_closed[of "H"], assumption+,
       simp)
done

section "6. preliminary lemmas for Zassenhaus"

lemma (in Group) Gp_sg_subset:"[|G » H; Gp G H » K|] ==> K ⊆ H"
by (frule Group_Gp[of "H"],
       frule Group.sg_subset[of "\<natural>H" "K"], assumption,
       thin_tac "(\<natural>H) » K", thin_tac "Group (\<natural>H)",
       simp add:Gp_def) 

lemma (in Group) inter_Gp_nsg:"[|G \<triangleright> N; G » H |] ==> (\<natural>H) \<triangleright> (H ∩ N)"
apply (frule Group_Gp[of "H"],
       rule Group.cond_nsg[of "Gp G H" "H ∩ N"], assumption+,
       frule nsg_sg[of "N"], frule inter_sgs[of "H" "N"], assumption+,
       rule sg_sg [of "H" "H ∩ N"], assumption+) 
 apply (rule subsetI, simp)
 apply ((rule ballI)+, simp, 
         simp add:Gp_carrier,
         simp add:Gp_mult_induced[of "H"],
         simp add:sg_i_induced[of "H"])
 apply (erule conjE,
        frule_tac x = a in sg_i_closed[of "H"], assumption+,
        frule_tac x = a and y = h in sg_mult_closed, assumption+,
        simp add:Gp_mult_induced[of "H"],
        simp add:sg_mult_closed)
 apply (frule_tac h = a in sg_subset_elem[of "H"], assumption+,
        simp add:nsgPr1_1[of "N"])
done

lemma (in Group) ZassenhausTr0:"[|G » H; G » H1; G » K; G » K1;
       Gp G H \<triangleright> H1; Gp G K \<triangleright> K1|] ==> Gp G (H ∩ K) \<triangleright> (H ∩ K1)"
apply (frule inter_sgs[of "H" "K"], assumption,
       frule inter_sgs[of "H" "K1"], assumption,
       frule Group_Gp[of "H"],
       frule Group_Gp[of "K"],
       frule Group.nsg_sg[of "\<natural>H" "H1"], assumption+,
       frule Group.nsg_sg[of "\<natural>K" "K1"], assumption+)
apply (rule Group.cond_nsg[of "\<natural>(H ∩ K)" "H ∩ K1"],
       simp add:Group_Gp[of "H ∩ K"])
apply (rule sg_sg[of "H ∩ K" "H ∩ K1"], assumption+)
apply (frule Gp_sg_subset[of "K" "K1"], assumption+,
       rule subsetI, simp add:subsetD)
apply ((rule ballI)+, simp) 

apply (frule Gp_sg_subset[of "K" "K1"], assumption+,
       erule conjE, frule_tac c = h in subsetD[of "K1" "K"], assumption+)
apply (rule conjI)
 apply (simp only:Gp_carrier,
        subst Gp_mult_induced1[of "H" "K"], assumption+, simp,
        simp only:sg_i_induced1)
 apply (frule_tac a = a and b = h in Group.mult_closed[of "\<natural>H"],
        simp add:Gp_carrier, simp add:Gp_carrier,
        simp only:Gp_carrier)
 apply (frule_tac a = a in Group.i_closed[of "\<natural>H"],
        simp add:Gp_carrier)
 apply (simp add:Gp_mult_induced1[of "H" "K"], simp add:Gp_carrier,
        subst Gp_mult_induced1[of "H" "K"], assumption+,
        simp add:Gp_mult_induced sg_mult_closed,
        simp add:sg_i_induced sg_i_closed) 
 apply (simp add:Gp_mult_induced sg_i_induced, simp add:sg_mult_closed)

 apply (subst Gp_mult_induced2[of "H" "K"], assumption+,
        simp add:Gp_carrier, simp, subst sg_i_induced2, assumption+,
        simp add:Gp_carrier)
 apply (frule_tac a = a and b = h in Group.mult_closed[of "\<natural>K"],
        simp add:Gp_carrier, simp add:Gp_carrier,
        frule_tac a = a in Group.i_closed[of "\<natural>K"], simp add:Gp_carrier)   
 apply (subst Gp_mult_induced2, assumption+,
        simp add:Gp_carrier, simp add:Gp_mult_induced sg_mult_closed,
        simp add:Gp_carrier, simp add:sg_i_induced sg_i_closed)
apply (rule_tac a = a and h = h in Group.nsgPr1_1[of "\<natural>K" "K1"], assumption+,
       simp add:Gp_carrier, assumption)
done 

lemma (in Group) lcs_sub_s_mult:"[|G » H; G » N; a ∈ H|] ==> a ♦ N ⊆ H \<struct>G N"
apply (rule subsetI)
apply (simp add:lcs_def s_top_def, blast)
done

lemma (in Group) rcs_sub_smult:"[|G » H; G » N; a ∈ H|] ==> N • a ⊆ N \<struct>G H"
apply (rule subsetI)
 apply (simp add:rcs_def s_top_def, blast)
done

lemma (in Group) smult_commute_sg_nsg:"[|G » H; G \<triangleright> N|] ==> H \<struct>G N = N \<struct>G H"
apply (frule nsg_sg[of "N"])
apply (rule equalityI)
 apply (rule subsetI,
        simp add:s_top_def, (erule bexE)+,
        frule_tac h = xa in sg_subset_elem, assumption+,
        frule_tac h = y in sg_subset_elem, assumption,
        frule_tac a = xa and b = y in mult_closed, assumption,
        frule_tac a = xa in i_closed,
        frule_tac a = "xa · y" and b = "ρ xa" and c = xa in tassoc,
        assumption+,
        frule sym, thin_tac "xa · y = x", simp,
        thin_tac "x = xa · y", simp add:l_i r_unit,
        frule_tac a = xa and h = y in nsgPr1_1[of "N"], assumption+)
 apply blast

 apply (rule subsetI)
 apply (simp add:s_top_def, (erule bexE)+,
        frule_tac h = xa in sg_subset_elem, assumption+,
        frule_tac h = y in sg_subset_elem, assumption,
        frule_tac a = xa and b = y in mult_closed, assumption,
        frule_tac a = y in i_closed,
        frule_tac a = y and b = "ρ y" and c = "xa · y" in tassoc,
        assumption+,
        frule sym, thin_tac "xa · y = x", simp,
        thin_tac "x = xa · y", simp add:r_i l_unit,
        frule_tac a = y and h = xa in nsgPr2[of "N"], assumption+,
        frule sym, thin_tac "xa · y = y · (ρ y · (xa · y))")
 apply blast
done  

lemma (in Group) smult_sg_nsg:"[|G » H; G \<triangleright> N|] ==> G » H \<struct>G N"
apply (frule  smult_commute_sg_nsg[of "H" "N"], assumption+,
       frule nsg_sg[of "N"],
       rule s_top_commute_sg[of "H" "N"], assumption+,
       rule sym, assumption)
done
         
lemma (in Group) smult_nsg_sg:"[|G » H; G \<triangleright> N|] ==> G » N \<struct>G H"
apply (frule smult_commute_sg_nsg[THEN sym, of "H" "N"], assumption+)
apply (simp add:smult_sg_nsg)
done

lemma (in Group) Gp_smult_sg_nsg:"[|G » H; G \<triangleright> N|] ==> Group (Gp G (H \<struct>G N))"
apply (frule smult_sg_nsg[of "H" "N"], assumption+)
apply (simp add:Group_Gp)
done  

lemma (in Group) N_sg_HN:"[|G » H; G \<triangleright> N|] ==> Gp G (H \<struct>G N) » N"
apply (frule smult_sg_nsg[of "H" "N"], assumption+,
       frule nsg_sg[of "N"],
       frule r_sub_smult[of "H" "N"], assumption+)
apply (rule sg_sg[of "H \<struct>G N" "N"], assumption+)
done

lemma (in Group) K_absorb_HK:"[|G » H; G » K; H ⊆ K|] ==>  H \<struct>G K = K"
apply (frule r_sub_smult[of "H" "K"], assumption+)
apply (rule equalityI)
apply (thin_tac "K ⊆ H \<struct>G K",
       rule subsetI, simp add:s_top_def, (erule bexE)+,
       frule_tac c = xa in subsetD[of "H" "K"], assumption+,
       frule_tac x = xa and y = y in sg_mult_closed[of "K"], assumption+,
       simp)
apply assumption
done
 
lemma (in Group) nsg_Gp_nsg:"[|G » H; G \<triangleright> N; N ⊆ H|] ==> Gp G H \<triangleright> N"
apply (frule Group_Gp[of "H"],
       frule nsg_sg[of "N"], 
       frule sg_sg[of "H" "N"], assumption+,
       rule Group.cond_nsg[of "\<natural>H" "N"], assumption+)
apply ((rule ballI)+,
       frule_tac c = h in subsetD[of "N" "H"], assumption+,
       simp add:Gp_carrier,
       simp add:Gp_mult_induced[of "H"] sg_i_induced[of "H"] 
                sg_mult_closed sg_i_closed)
apply (rule_tac a = a and h = h in nsgPr1_1[of "N"], assumption+,
       rule_tac h = a in sg_subset_elem[of "H"], assumption+)
done

lemma (in Group) Gp_smult_nsg:"[|G » H; G \<triangleright> N|] ==> Gp G (H \<struct>G N) \<triangleright> N"
apply (frule smult_sg_nsg[of "H" "N"], assumption+,
       frule nsg_sg[of "N"],
       frule N_sg_HN[of "H" "N"], assumption+,
       frule Gp_smult_sg_nsg[of "H" "N"], assumption+,
       rule Group.cond_nsg[of "\<natural>(H \<struct>G N)" "N"], assumption+)
apply ((rule ballI)+,
       frule_tac a = a in Group.i_closed[of "\<natural>(H \<struct>G N)"], assumption+,
       simp add:Gp_carrier) 

apply (frule r_sub_smult[of "H" "N"], assumption+,
       frule_tac c = h in subsetD[of "N" "H \<struct>G N"], assumption+,
       simp add:Gp_mult_induced[of "H \<struct>G N"] sg_i_induced[of "H \<struct>G N"])
apply (frule_tac x = a and y = h in sg_mult_closed[of "H \<struct>G N"], assumption+,
       simp add:Gp_mult_induced)
apply (rule_tac a = a and h = h in nsgPr1_1[of "N"], assumption+,
       frule sg_subset[of "H \<struct>G N"], frule_tac c = a in subsetD[of "H \<struct>G N"
       "carrier G"], assumption+)
done 

lemma (in Group) Gp_smult_nsg1:"[|G » H; G \<triangleright> N|] ==> Gp G (N \<struct>G H) \<triangleright> N"
apply (simp add:smult_commute_sg_nsg[THEN sym, of "H" "N"],
       simp only:Gp_smult_nsg)
done

lemma (in Group) ZassenhausTr2_3:"[|G » H; G » H1; Gp G H \<triangleright> H1|] ==> H1 ⊆ H"
apply (frule Group_Gp[of "H"],
       frule Group.nsg_sg[of "\<natural>H" "H1"], assumption,
       frule Group.sg_subset[of "\<natural>H" "H1"], assumption, simp add:Gp_carrier)
done

lemma (in Group) ZassenhausTr2_4:"[|G » H; G » H1; Gp G H \<triangleright> H1; h ∈ H; 
               h1 ∈ H1|] ==> h · h1 · (ρ h) ∈ H1" 
apply (frule Group_Gp[of "H"])
apply (frule_tac a = h and h = h1 in Group.nsgPr1_1[of "\<natural>H" "H1"], assumption)
       apply (simp add:Gp_carrier) apply assumption
apply (simp add:Gp_def)
done

lemma (in Group) ZassenhausTr1:"[|G » H; G » H1; G » K; G » K1;
       Gp G H \<triangleright> H1; Gp G K \<triangleright> K1|] ==> H1 \<struct>G (H ∩ K1) = (H ∩ K1) \<struct>G H1"
apply (frule Group_Gp[of "H"],
       frule Group.nsg_sg[of "\<natural>H" "H1"], assumption,
       frule Group.sg_subset[of "\<natural>H" "H1"], assumption, simp add:Gp_carrier)
apply (frule Group_Gp[of "K"],
       frule Group.nsg_sg[of "\<natural>K" "K1"], assumption,
       frule Group.sg_subset[of "\<natural>K" "K1"], assumption, simp add:Gp_carrier)
apply (rule equalityI)
 apply (rule subsetI,
        simp add:s_top_def, (erule bexE)+,
        frule_tac h = xa in sg_subset_elem[of "H1"], assumption+,
        frule_tac h = y in sg_subset_elem[of "H"], simp,
        frule_tac a = y in i_closed,
        frule_tac a = xa and b = y in mult_closed, assumption+,
        frule_tac a1 = y and b1 = "ρ y" and c1 = "xa · y" in tassoc[THEN sym],
        assumption+)
 apply (frule sym, thin_tac "xa · y = x", simp, thin_tac "x = xa · y",
        simp add:r_i l_unit,
        frule_tac x = y in sg_i_closed[of "H"], simp) 
 apply (frule_tac a1 = "ρ y" and b1 = xa and c1 = y in tassoc[THEN sym],
         assumption+, simp, thin_tac "ρ y · (xa · y) = ρ y · xa · y")
 apply (frule_tac h = "ρ y" and ?h1.0 = xa in ZassenhausTr2_4[of "H" "H1"],
                assumption+, simp add:iop_i_i)
 apply blast

 apply (rule subsetI)
 apply (simp add:s_top_def, (erule bexE)+,
        frule_tac h = xa in sg_subset_elem[of "H"], simp,
        frule_tac h = y in sg_subset_elem[of "H1"], assumption,
        frule sym, thin_tac "xa · y = x", simp, thin_tac "x = xa · y")
apply (frule_tac a = xa in i_closed,
       frule_tac a = xa and b = y in mult_closed, assumption+,
       frule_tac a = "xa · y" and b = "ρ xa" and c = xa in tassoc, assumption+)
 apply (simp add:l_i r_unit,
       frule_tac h = xa and ?h1.0 = y in ZassenhausTr2_4[of "H" "H1"],
                assumption+, simp, assumption, blast)
done

lemma (in Group) ZassenhausTr1_1:"[|G » H; G » H1; G » K; G » K1;
       Gp G H \<triangleright> H1; Gp G K \<triangleright> K1|] ==> G » (H1 \<struct>G (H ∩ K1))"
apply (rule s_top_commute_sg, assumption)
apply (simp add:inter_sgs[of "H" "K1"])
apply (rule ZassenhausTr1 [THEN sym, of "H" "H1" "K" "K1"], assumption+)
done

lemma (in Group) ZassenhausTr2:"[|G » H; G » H1; G » K; Gp G H \<triangleright> H1|] ==>
                                  H1 \<struct>G (H ∩ K) = (H ∩ K) \<struct>G H1"
apply (frule special_nsg_G1[of "K"])
apply (simp add: ZassenhausTr1 [of "H" "H1" "K" "K"])
done

lemma (in Group) ZassenhausTr2_1:"[|G » H; G » H1; G » K; Gp G H \<triangleright> H1|]
  ==> G » H1 \<struct>G (H ∩ K)"
apply (frule ZassenhausTr2 [of "H" "H1" "K"], assumption+,
       frule inter_sgs [of "H" "K"], assumption+,
       rule s_top_commute_sg, assumption+)
apply (rule sym, assumption)
done

lemma (in Group) ZassenhausTr2_2:"[|G » H; G » H1; G » K; G » K1;
       Gp G H \<triangleright> H1; Gp G K \<triangleright> K1|] ==> H1 \<struct>G (H ∩ K1) ⊆  H1 \<struct>G (H ∩ K)"  
apply (frule Group_Gp[of "K"],
       frule Group.nsg_sg[of "\<natural>K" "K1"], assumption,
       frule Group.sg_subset[of "\<natural>K" "K1"], assumption, simp add:Gp_carrier,
       frule Group_Gp[of "H"],
       frule Group.nsg_sg[of "\<natural>H" "H1"], assumption,
       frule Group.sg_subset[of "\<natural>H" "H1"], assumption, simp add:Gp_carrier,
       frule sg_subset[of "H"], frule sg_subset[of "K"])
apply (rule s_top_mono[of "H1" "H ∩ K" "H1" "H ∩ K1"],
       rule subset_trans[of "H1" "H" "carrier G"], assumption+,
       blast, simp, blast)
done

lemma (in Group) ZassenhausTr2_5:"[|G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
      Gp G K \<triangleright> K1; a∈ H1; b ∈ H ∩ K1; c ∈ H1|] ==>
                      a · b · c ∈ H1 \<struct>G (H ∩ K1)"
apply (simp, erule conjE)
apply (frule sg_subset_elem[of "H1" "a"], assumption+, 
       frule sg_subset_elem[of "H1" "c"], assumption+,
       frule sg_subset_elem[of "H" "b"], assumption+,
       frule i_closed[of "b"],
       frule mult_closed[of "a" "b"], assumption+,
       frule mult_closed[of "a · b" "c"], assumption+,
       frule tassoc[of "a · b · c" "ρ b" "b"], assumption+,
       simp add:l_i r_unit)
apply (rule eq_elem_in[of "a · b · c · ρ b · b" "H1 \<struct>G H ∩ K1" "a · b · c"],
       thin_tac "a · b · c · ρ b · b = a · b · c",
       frule inter_sgs[of "H" "K1"], assumption+,
       rule s_top_elem[of "H1" "H ∩ K1" "a · b · c · ρ b " "b"], assumption+,
       subst tOp_assocTr42, assumption+,
       frule mult_closed[of "b" "c"], assumption+,
       simp add:tassoc[of "a" "b · c" "ρ b"])
apply (rule sg_mult_closed[of "H1" "a" "b · c · ρ b"], assumption+,
       rule ZassenhausTr2_4[of "H" "H1" "b" "c"], assumption+) apply blast
apply assumption
done

lemma (in Group) ZassenhausTr2_6:"[|u ∈ carrier G; v ∈ carrier G;  
       x ∈ carrier G; y ∈ carrier G|] ==> 
       (u · v) · (x · y) · (ρ (u · v)) = 
                          u · v · x · (ρ v) · (v · y · (ρ v)) · (ρ u)"
apply (simp add:i_ab)
apply (frule i_closed[of "u"], frule i_closed[of "v"])
apply (frule mult_closed[of "u" "v"], assumption+,
       frule mult_closed[of "u · v" "x"], assumption+,
       frule mult_closed[of "v" "y"], assumption+,
       frule mult_closed[of "v · y" "ρ v"], assumption+,
       frule mult_closed[of "u · v · x" "ρ v"], assumption+,
       simp add:tOp_assocTr42[THEN sym, of "u · v · x · ρ v " 
                               "v · y" "ρ v" "ρ u"])
apply (frule mult_closed[of "x" "y"], assumption+,
       frule mult_closed[of "u · v" "x · y"], assumption+)
apply (simp add:tassoc[THEN sym, of "u · v · (x · y)" "ρ v" "ρ u"])
apply (rule r_mult_eqn[of _ _ "ρ u"],
       rule mult_closed[of "u · v · (x · y)" "ρ v"], assumption+,
       (rule mult_closed)+, assumption+)
apply (rule r_mult_eqn[of _ _ "ρ v"], assumption+,
       (rule mult_closed)+, assumption+)
apply (simp add:tOp_assocTr41[THEN sym, of "u · v · x" "ρ v" "v" "y"],
       simp add:tOp_assocTr42[of "u · v · x" "ρ v" "v" "y"],
       simp add:l_i r_unit)
apply (simp add:tOp_assocTr41)
done

lemma (in Group) ZassenhausTr2_7:"[|a ∈ carrier G; x ∈ carrier G; y ∈ carrier G|]
        ==> a · ( x · y) · (ρ a) = a · x · (ρ a) · (a · y · (ρ a))"
apply (frule i_closed[of "a"], 
       frule mult_closed[of "a" "y"], assumption+,
       frule mult_closed[of "a · y" "ρ a"], assumption+)
apply (simp add:tOp_assocTr41[of "a" "x" "ρ a" "a · y · (ρ a)"],
       simp add:tassoc[THEN sym, of "ρ a" "a · y" "ρ a"],
       simp add:tassoc[THEN sym, of "ρ a" "a" "y"] l_i l_unit,
       simp add:tOp_assocTr41[THEN sym],
       simp add:tOp_assocTr42[THEN sym])
done

lemma (in Group) ZassenhausTr3:"[|G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
              Gp G K \<triangleright> K1|] ==>  Gp G (H1 \<struct>G (H ∩ K)) \<triangleright> (H1 \<struct>G (H ∩ K1))"
apply (frule ZassenhausTr2_1 [of "H" "H1" "K"], assumption+,
       frule ZassenhausTr2_1 [of "H" "H1" "K1"], assumption+,
       frule ZassenhausTr2_2 [of "H" "H1" "K" "K1"], assumption+,
       frule sg_sg [of "H1 \<struct>G (H ∩ K)" "H1 \<struct>G (H ∩ K1)"], assumption+,
       frule Group_Gp[of "H1 \<struct>G H ∩ K"])
apply (rule Group.cond_nsg[of "\<natural>(H1 \<struct>G H ∩ K)" "H1 \<struct>G H ∩ K1"], assumption+,
       (rule ballI)+,
       simp add:Gp_carrier)
 apply (frule_tac c = h in subsetD[of "H1 \<struct>G H ∩ K1" "H1 \<struct>G H ∩ K"],
            assumption+,
        simp add:Gp_mult_induced[of "H1 \<struct>G H ∩ K"],
        simp add:sg_i_induced[of "H1 \<struct>G H ∩ K"],
        frule_tac x = a in sg_i_closed[of "H1 \<struct>G H ∩ K"], assumption+,
        frule_tac x = a and y = h in sg_mult_closed[of "H1 \<struct>G H ∩ K"], 
        assumption+,
        simp add:Gp_mult_induced[of "H1 \<struct>G H ∩ K"],
        thin_tac "ρ a ∈ H1 \<struct>G H ∩ K", thin_tac "a · h ∈ H1 \<struct>G H ∩ K")
 apply (simp add:s_top_def[of "G" "H1" "H ∩ K"], (erule bexE)+,
        simp add:s_top_def[of "G" "H1" "H ∩ K1"], fold s_top_def,
       (erule bexE)+, thin_tac "xa · ya = h", (erule conjE)+,
        thin_tac "xa ∈ H1", thin_tac "ya ∈ H",
        frule sym, thin_tac "x · y = a", 
        frule sym, thin_tac "xb · yb = h", simp, (erule conjE)+,
        thin_tac "a = x · y", thin_tac "h = xb · yb")
 apply (frule_tac h = x in sg_subset_elem[of "H1"], assumption+,
        frule_tac h = y in sg_subset_elem[of "H"], assumption+,
        frule_tac h = xb in sg_subset_elem[of "H1"], assumption+,
        frule_tac h = yb in sg_subset_elem[of "H"], assumption+,
        subst ZassenhausTr2_6, assumption+)

apply (frule_tac a = y and b = xb in mult_closed, assumption+,
       frule_tac a = y in i_closed,
       frule_tac a = "y · xb" and b = "ρ y" in mult_closed, assumption+,
       frule_tac a = x and b = "y · xb · ρ y" in mult_closed, assumption+,
       frule_tac a = y and b = yb in mult_closed, assumption+,
       frule_tac a = "y · yb" and b = "ρ y" in mult_closed, assumption+,
       frule_tac a = "x · y · xb · ρ y" and b = "y · yb · ρ y" and c = "ρ x" 
        in ZassenhausTr2_5[of "H" "H1" "K" "K1"], assumption+,
       frule_tac a = x and b = y in mult_closed, assumption+,
       frule_tac a = "x · y" and b = xb and c = "ρ y" in tassoc, assumption+,
       simp, thin_tac "x · y · xb · ρ y = x · y · (xb · ρ y)",
       frule_tac a = xb and b = "ρ y" in mult_closed, assumption+,
              simp add:tassoc)
apply (rule_tac x = x and y = "y · (xb · ρ y)" in sg_mult_closed, assumption+,
       simp add:tassoc[THEN sym],
       rule_tac h = y and ?h1.0 = xb in ZassenhausTr2_4[of "H" "H1"],
            assumption+)
apply (frule_tac x = y and y = yb in sg_mult_closed[of "H"], assumption+,
       frule_tac x = y in sg_i_closed[of "H"], assumption+,
       frule_tac x = "y · yb" and y = "ρ y" in sg_mult_closed[of "H"], 
       assumption+, simp,
       rule_tac h = y and ?h1.0 = yb in ZassenhausTr2_4[of "K" "K1"],
            assumption+,
       rule_tac x = x in sg_i_closed[of "H1"], assumption+)
apply (simp add:s_top_def[of "G" "H1" "H ∩ K1"])
done

lemma (in Group) ZassenhausTr3_2:"[|G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
              Gp G K \<triangleright> K1|] ==> G » H1 \<struct>G (H ∩ K1) \<struct>G (H ∩ K)" 
apply (frule s_top_assoc[of "H1" "H ∩ K1" "H ∩ K"],
       (simp add:inter_sgs)+,
       frule inter_sgs[of "H" "K1"], assumption+,
       frule inter_sgs[of "H" "K"], assumption+)
 apply (frule K_absorb_HK[of "H ∩ K1" "H ∩ K"], assumption+,
        frule ZassenhausTr2_3[of "K" "K1"], assumption+, blast,
        simp, simp add:ZassenhausTr2_1)
done

lemma (in Group) ZassenhausTr3_3:"[|G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1|] ==> (H1 ∩ K) \<struct>G (H ∩ K1) = (K1 ∩ H) \<struct>G (K ∩ H1)"
apply (rule equalityI)
 apply (rule subsetI, simp add:s_top_def, (erule bexE)+)
 apply (frule sym, thin_tac "xa · y = x", simp, (erule conjE)+)
 apply (frule_tac h = xa in sg_subset_elem[of "H1"], assumption+,
        frule_tac h = y in sg_subset_elem[of "K1"], assumption+,
        frule_tac a = xa in i_closed,
        frule_tac a = xa and b = y and c = "ρ xa" and d = xa in tOp_assocTr41,
        assumption+,
        frule_tac a = xa and b = y in mult_closed, assumption,
        simp add:l_i r_unit)
 apply (frule_tac h = xa and ?h1.0 = y in ZassenhausTr2_4[of "K" "K1"],
        assumption+)
 apply (frule ZassenhausTr2_3[of "H" "H1"], assumption+,
        frule_tac c = xa in subsetD[of "H1" "H"], assumption+)
 apply (frule_tac x = xa and y = y in sg_mult_closed[of "H"], assumption+)
 apply (frule_tac x = xa in sg_i_closed[of "H"], assumption+,
        frule_tac x = "xa · y" and y = "ρ xa" in sg_mult_closed[of "H"],
                       assumption+)
 apply blast

 apply (rule subsetI, simp add:s_top_def, (erule bexE)+)
 apply (frule sym, thin_tac "xa · y = x", simp, (erule conjE)+)
 apply (frule_tac h = xa in sg_subset_elem[of "K1"], assumption+,
        frule_tac h = y in sg_subset_elem[of "H1"], assumption+,
        frule_tac a = xa in i_closed,
        frule_tac a = xa and b = y and c = "ρ xa" and d = xa in tOp_assocTr41,
        assumption+,
        frule_tac a = xa and b = y in mult_closed, assumption,
        simp add:l_i r_unit)

 apply (frule_tac h = xa and ?h1.0 = y in ZassenhausTr2_4[of "H" "H1"],
        assumption+)
 apply (frule ZassenhausTr2_3[of "K" "K1"], assumption+)
 apply (frule_tac c = xa in subsetD[of "K1" "K"], assumption+)
 apply (frule_tac x = xa and y = y in sg_mult_closed[of "K"], assumption+)
 apply (frule_tac x = xa in sg_i_closed[of "K"], assumption+,
        frule_tac x = "xa · y" and y = "ρ xa" in sg_mult_closed[of "K"],
                       assumption+)
 apply blast
done

lemma (in Group) ZassenhausTr3_4:"[|G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1; g ∈ H ∩ K; h ∈ H ∩ K1|] ==> g · h · (ρ g) ∈ H ∩ K1"
apply (simp, (erule conjE)+) 
apply (frule_tac x = g and y = h in sg_mult_closed, assumption+,
       frule_tac x = g in sg_i_closed[of "H"], assumption+,
       simp add:sg_mult_closed[of "H" "g · h" "ρ g"])
apply (rule ZassenhausTr2_4[of "K" "K1" "g" "h"], assumption+)
done

lemma (in Group) ZassenhausTr3_5:"[|G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1|] ==> (Gp G (H ∩ K)) \<triangleright> (H1 ∩ K) \<struct>G (H ∩ K1)"
apply (frule inter_sgs[of "H" "K"], assumption,
       frule inter_sgs[of "H1" "K"], assumption,
       frule inter_sgs[of "K" "H"], assumption,
       frule inter_sgs[of "H" "K1"], assumption+)
apply (frule ZassenhausTr3[of "H ∩ K" "H1 ∩ K" "K ∩ H" "H ∩ K1"],
       assumption+,
       frule ZassenhausTr0[of "K" "K1" "H" "H1"], assumption+,
       simp add:Int_commute,
       frule ZassenhausTr0[of "H" "H1" "K" "K1"], assumption+,
       simp add:Int_commute)
 apply (frule ZassenhausTr2_3 [of "K" "K1"], assumption+,
        frule ZassenhausTr2_3 [of "H" "H1"], assumption+) 
 apply (simp add:Int_commute[of "K" "H"])
 apply (cut_tac Int_mono[of "H" "H" "K1" "K"])
 apply (cut_tac Int_mono[of "H1" "H" "K" "K"])
 apply (simp only:Int_absorb1[of "H ∩ K1" "H ∩ K"],
        simp only:K_absorb_HK[of "H1 ∩ K" "H ∩ K"]) apply simp+
done

lemma (in Group) ZassenhausTr4:"[|G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
     Gp G K \<triangleright> K1|] ==> (H1 \<struct>G (H ∩ K1)) \<struct>G (H1 \<struct>G (H ∩ K)) = H1 \<struct>G (H ∩ K)"
apply (frule ZassenhausTr2 [of "H" "H1" "K"], assumption+,
       frule ZassenhausTr2 [of "H" "H1" "K1"], assumption+,
       frule ZassenhausTr1_1 [of "H" "H1" "K" "K1"], assumption+,
       frule ZassenhausTr2_1 [of "H" "H1" "K"], assumption+,
       frule ZassenhausTr2_2 [of "H" "H1" "K" "K1"], assumption+)
apply (rule K_absorb_HK[of "H1 \<struct>G H ∩ K1" "H1 \<struct>G H ∩ K"], assumption+)
done

lemma (in Group) ZassenhausTr4_0: "[|G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
     Gp G K \<triangleright> K1|] ==>  H1 \<struct>G (H ∩ K) = (H1 \<struct>G (H ∩ K1)) \<struct>G (H ∩ K)"
apply (frule inter_sgs [of "H" "K1"], assumption+,
       frule inter_sgs [of "H" "K"], assumption+)
apply (subst s_top_assoc [of "H1" "H ∩ K1" "H ∩ K"], assumption+,
       subst K_absorb_HK[of "H ∩ K1" "H ∩ K"], assumption+)
apply (frule ZassenhausTr2_3[of "K" "K1"], assumption+,
       rule Int_mono[of "H" "H" "K1" "K"]) 
apply simp+
done

lemma (in Group) ZassenhausTr4_1:"[|G » H; (Gp G H) \<triangleright> H1; (Gp G H) » (H ∩ K)|]
                           ==> (Gp G (H1 \<struct>G (H ∩ K))) \<triangleright> H1" 
apply (frule Group_Gp [of "H"],
       frule Group.nsg_sg[of "Gp G H" "H1"], assumption+,
       frule Group.Gp_smult_nsg1[of "\<natural>H" "H ∩ K" "H1"], assumption+,
       frule subg_sg_sg [of "H" "H1"], assumption+,
       frule Group.sg_subset[of "\<natural>H" "H1"], assumption,
       frule Group.sg_subset[of "\<natural>H" "H ∩ K"], assumption+,
       frule Group.smult_nsg_sg[of "\<natural>H" "H ∩ K" "H1"], assumption+,
       frule Group.s_top_sub[of "\<natural>H" "H1" "H ∩ K"], assumption+,
       simp add:Gp_carrier,
       simp only:s_top_induced[of "H" "H1" "H ∩ K"])
apply (frule subg_sg_sg[of "H" "H1 \<struct>G H ∩ K"], assumption+,
       simp add:Gp_inherited[of "H1 \<struct>G H ∩ K" "H"])
done

section "7. homomorphism"

lemma gHom: "[|Group F; Group G; f ∈ gHom F G ; x ∈ carrier F; 
           y ∈ carrier F|]  ==> f (x ·F y) = (f x) ·G (f y)"
apply (simp add: gHom_def)
done

lemma gHom_mem:"[|Group F; Group G; f ∈ gHom F G ; x ∈ carrier F|] ==>
                             (f x) ∈ carrier G" 
apply (simp add:gHom_def, (erule conjE)+)
apply (simp add:funcset_mem)
done

lemma gHom_func:"[|Group F; Group G; f ∈ gHom F G|] ==>
                            f ∈ carrier F -> carrier G"
by (simp add:gHom_def)

(* we have to check the composition of two ghom's *) 
lemma gHomcomp:"[|Group F; Group G; Group H; f ∈ gHom F G; g ∈ gHom G H|] 
    ==> (g oF f) ∈ gHom F H"
apply (simp add:gHom_def)
 apply (erule conjE)+
apply (simp add:cmpghom_def composition)
apply (rule ballI)+
 apply (simp add:compose_def)
 apply (frule_tac x = x in funcset_mem [of "f" "carrier F" "carrier G"],
                                              assumption+) 
apply (frule_tac x = y in funcset_mem [of "f" "carrier F" "carrier G"],
                                              assumption+)
apply (simp add:Group.mult_closed[of "F"]) 
done

lemma gHom_comp_gsurjec:"[|Group F; Group G; Group H; gsurjF,G f; 
  gsurjG,H g|]  ==> gsurjF,H (g oF f)" 
apply (simp add:gsurjec_def,
       simp add:gHomcomp,
       (erule conjE)+)
 apply (simp add:cmpghom_def,
        simp add:gHom_def, (erule conjE)+,
        rule compose_surj, assumption+)
done

lemma gHom_comp_ginjec:"[|Group F; Group G; Group H; ginjF,G f; ginjG,H g|] ==> 
                          ginjF,H (g oF f)" 
apply (simp add:ginjec_def,
       simp add:gHomcomp,
       simp add:gHom_def,
       (erule conjE)+)
apply (simp add:cmpghom_def,
       simp add:comp_inj)
done

lemma ghom_unit_unit:"[|Group F; Group G; f ∈ gHom F G |] ==>
                                                   f (\<one>F) = \<one>G"
apply (frule Group.unit_closed[of "F"])
apply (frule gHom [of "F" "G" "f" "\<one>F" "\<one>F"], assumption+)
 apply (simp add:Group.l_unit[of "F"])
 apply (frule gHom_mem[of "F" "G" "f" "\<one>F"], assumption+)
 apply (frule sym)
 apply (rule Group.one_unique[of "G" "f \<one>F" "f \<one>F"], assumption+)
done

lemma ghom_inv_inv:"[|Group F; Group G; f ∈ gHom F G ; x ∈ carrier F|] ==>
   f (ρF x) = ρG (f x)"
apply (frule Group.i_closed[of "F" "x"], assumption+,
       frule gHom [of "F" "G" "f" "ρF x" "x"], assumption+)
apply (simp add:Group.l_i, simp add:ghom_unit_unit)
apply (frule sym,
       frule gHom_mem[of "F" "G" "f" "x"], assumption+ ,
       frule gHom_mem[of "F" "G" "f" "ρF x"], assumption+,
       rule Group.l_i_unique[THEN sym, of "G" "f x" "f (ρF x)"], assumption+)
done

lemma ghomTr3:"[|Group F; Group G; f ∈ gHom F G ; x ∈ carrier F;
       y ∈ carrier F; f (x ·FF y)) = \<one>G |] ==> f x = f y"
apply (frule Group.i_closed[of "F" "y"], assumption+,
       simp only:gHom ghom_inv_inv)
apply (rule Group.r_div_eq[of "G" "f x" "f y"], assumption,
       (simp add:gHom_mem)+)
done

lemma iim_nonempty:"[|Group F; Group G; f ∈ gHom F G; G » K|] ==>
                       (iim F G f K) ≠ {}" 
apply (frule Group.sg_unit_closed[of "G" "K"], assumption+,
       frule Group.unit_closed[of "F"])
apply (frule ghom_unit_unit[of "F" "G" "f"], assumption+, simp add:iim_def,
       frule sym, thin_tac "f \<one>F = \<one>G", simp)
apply blast
done

lemma ghomTr4:"[|Group F; Group G; f ∈ gHom F G; G » K|] ==> 
                  F » (iim F G f K)"
apply (rule Group.sg_condition[of "F" "iim F G f K"], assumption+,
       rule subsetI, simp add:iim_def,
       simp add:iim_nonempty)
apply ((rule allI)+, rule impI, erule conjE)
apply (simp add:iim_def) apply (erule conjE)+
apply (frule_tac a = b in Group.i_closed[of "F"], assumption+,
       frule_tac a = a and b = "ρF b" in Group.mult_closed[of "F"],
        assumption+, simp)
apply (simp add:gHom ghom_inv_inv)
apply (frule_tac x = "f b" in Group.sg_i_closed[of "G" "K"], assumption+) 
apply (simp add:gHom_mem Group.sg_mult_closed)
done

lemma (in Group) IdTr0: "idmap (carrier G) ∈ gHom G G"
apply (simp add:gHom_def)
apply (simp add:idmap_def extensional_def)
apply (simp add:Pi_def mult_closed)
done

syntax 
  "@IDMAP"::"('a, 'more) Group_scheme => ('a => 'a)"
                ("(I_)" [999]1000)

  "@INVFUN"::"[('a, 'm1) Group_scheme, ('b, 'm) Group_scheme, 'a => 'b]
              => ('b => 'a)" ("(3Ifn _ _ _)" [88,88,89]88)
translations
   "IF" == "idmap (carrier F)"
   "Ifn F G f" == "invfun (carrier F) (carrier G) f"

lemma IdTr1:"[|Group F; x ∈ carrier F|] ==> (IF) x = x"
apply (simp add:idmap_def)
done

lemma IdTr2:"Group F ==> gbijF,F (IF)"
apply (simp add:gbijec_def)
apply (rule conjI)
 (* gsurjec *)
  apply (simp add:gsurjec_def)
      apply (rule conjI)
        apply (simp add:idmap_def gHom_def Group.mult_closed)
      apply (rule univar_func_test, simp add:idmap_def)
      apply (simp add:surj_to_def)
      apply (simp add:image_def idmap_def)
  (* ginjec *)
  apply (simp add:ginjec_def)
     apply (simp add:idmap_def gHom_def Group.mult_closed)
     apply (rule univar_func_test, simp add:idmap_def)
done

lemma Id_l_unit:"[|Group G; gbijG,G f|] ==> IG oG f = f"
apply (rule funcset_eq[of _ "carrier G"])
 apply (simp add:cmpghom_def)
 apply (simp add:gbijec_def gsurjec_def gHom_def)
apply (rule ballI)
 apply (simp add:cmpghom_def compose_def)
 apply (frule_tac x = x in gHom_mem[of "G" "G" "f"], assumption+)
  apply (simp add:gbijec_def gsurjec_def, assumption)
 apply (simp add:IdTr1)
done


section "8. gkernel"

lemma gkernTr1:"[|Group F; Group G; f ∈ gHom F G; x ∈ gkerF,G f|] ==>                    x ∈ carrier F"
apply (simp add: gkernel_def)
done 

lemma gkernTr1_1:"[|Group F; Group G; f ∈ gHom F G|] ==> gkerF,G f ⊆ carrier F"
by (rule subsetI, simp add:gkernTr1)

lemma gkernTr2:"[|Group F; Group G; f ∈ gHom F G; x ∈ gkerF,G f; y ∈ gkerF,G f|]
   ==> (x ·F y) ∈ gkerF,G f"
apply (simp add:gkernel_def)
apply (simp add:gHom, (erule conjE)+)
apply (simp add:Group.mult_closed, 
       frule Group.unit_closed[of "G"], 
       simp add:Group.l_unit[of "G"])
done

lemma gkernTr3:"[|Group F; Group G; f ∈ gHom F G ; x ∈ gkerF,G f|] ==>
                    (ρF x) ∈ gkerF,G f"
apply (simp add:gkernel_def)
(* thm ghom_inv_inv [of "F" "G" "f" "x"] *)
apply (simp add:ghom_inv_inv [of "F" "G" "f" "x"])
apply (simp add:Group.i_closed) apply (simp add:Group.i_one)
done

lemma gkernTr6:"[|Group F; Group G; f ∈ gHom F G|] ==> (\<one>F) ∈ gkerF,G f"
apply (simp add:gkernel_def)   
apply (simp add:Group.unit_closed ghom_unit_unit [of "F" "G" "f" ])
done

lemma gkernTr7:"[|Group F; Group G; f ∈ gHom F G |] ==> F » gkerF,G f"
apply (rule Group.sg_condition[of "F" "gkerF,G f"], assumption+,
       rule subsetI, simp add:gkernTr1,
       frule gkernTr6[of "F" "G" "f"], assumption+, simp add:nonempty)

apply ((rule allI)+, rule impI, erule conjE,
        frule_tac x = b in gkernTr3[of "F" "G" "f"], assumption+)
apply (simp add:gkernTr2)
done

lemma gker_normal:"[|Group F; Group G; f ∈ gHom F G|] ==> F \<triangleright> gkerF,G f"
apply (rule Group.cond_nsg[of "F" "gkerF,G f"], assumption)
apply (simp add:gkernTr7)
apply (rule ballI)+
apply (simp add:gkernel_def, erule conjE)
apply (frule_tac a = a in Group.i_closed[of "F"], assumption)
apply (subst gHom [of "F" "G" "f" _], assumption+)
       apply (simp add:Group.mult_closed[of "F"])
       apply assumption
apply (simp add:gHom)
 apply (simp add:Group.mult_closed[of "F"])+
 apply (frule_tac x = a in gHom_mem[of "F" "G" "f"], assumption+,
        simp add:Group.r_unit[of "G"]) 
 apply (simp add:ghom_inv_inv, simp add:Group.r_i)
done

lemma Group_coim:"[|Group F; Group G; f ∈ gHom F G|] ==> Group ( F/ gkerF,G f)"
by (frule gker_normal[of "F" "G" "f"], assumption+,
       simp add:Group.Group_Qg[of "F" "gkerF,G f"])

lemma gkern1:"[|Group F; Ugp E; f ∈ gHom F E|] ==> gkerF,E f = carrier F"
apply (frule Group_Ugp[of "E"])
apply (frule gkernTr1_1[of "F" "E" "f"], assumption+)
apply (rule equalityI, assumption)
apply (rule subsetI,
       thin_tac "gkerF,E f ⊆ carrier F",
       simp add:gkernel_def,
       frule_tac x = x in gHom_mem[of "F" "E" "f"], assumption+,
       simp add:Ugp_def)
done

lemma gkern2:"[|Group F; Group G; f ∈ gHom F G; ginjF,G f|] ==>
               gkerF,G f = {\<one>F}"
apply (frule gkernTr6[of "F" "G" "f"], assumption+,
       frule singleton_sub[of "\<one>F" "gkerF,G f"],
       rule equalityI)
 apply (rule subsetI,
        thin_tac "{\<one>F} ⊆ gkerF,G f",
        simp add:gkernel_def, (erule conjE)+)
 apply (frule sym, thin_tac "f \<one>F = \<one>G", simp, thin_tac "\<one>G = f \<one>F",
        simp add:ginjec_def, simp add:inj_on_def)  
apply assumption
done

lemma gkernTr9:"[|Group F; Group G; f ∈ gHom F G; a ∈ carrier F; b ∈ carrier F|]
             ==>  ((gkerF,G f) •F a = (gkerF,G f) •F b) = (f a = f b)"
(* thm r_cosEqVar1 [of "F" "ker f (F \<rightharpoonup> G) " "a" "b"] *)
apply (frule gkernTr7[of "F" "G" "f"], assumption+)
apply (subst Group.rcs_eq[THEN sym, of "F" "gkerF,G f" "a" "b"], assumption+)
apply (thin_tac "F » gkerF,G f")
apply (simp add:gkernel_def)
apply (frule Group.i_closed[of "F" "a"], assumption)
       apply (simp add:Group.mult_closed[of "F"])
       apply (simp add:gHom, simp add:ghom_inv_inv)
apply (frule gHom_mem[of "F" "G" "f" "a"], assumption+,
       frule gHom_mem[of "F" "G" "f" "b"], assumption+)
apply (rule iffI)
apply (rule Group.r_div_eq[THEN sym, of "G" "f b" "f a"], assumption+)
apply (simp add:Group.r_i[of "G"])
done 

lemma gkernTr11:"[|Group F; Group G; f ∈ gHom F G ; a ∈ carrier F|] ==> 
                  (iim F G f {f a}) = (gkerF,G f) •F a"
apply (frule gkernTr7[of "F" "G" "f"], assumption+)
apply (rule equalityI) (** iim F G f {f a} ⊆ kerF,Gf F a **)
 apply (rule subsetI,
        simp add:iim_def,
        erule conjE)
 apply (frule_tac a1 = x in gkernTr9[THEN sym, of "F" "G" "f" _ "a"],
            assumption+, simp,
        frule_tac a = x in Group.a_in_rcs[of "F" "gkerF,G f"], assumption+,
        simp)

 apply (rule subsetI)
 apply (simp add:gkernel_def rcs_def iim_def, erule exE, (erule conjE)+,
        rotate_tac -1, frule sym, thin_tac "h ·F a = x", simp add:gHom,
        thin_tac "x = h ·F a",
        frule gHom_mem[of "F" "G" "f" "a"], assumption+,
        simp add:Group.mult_closed Group.l_unit)
done

lemma gbij_comp_bij:"[|Group F; Group G; Group H; gbijF,G f; gbijG,H g|]
                   ==> gbijF,H (g oF f)"
apply (simp add:gbijec_def)
apply (simp add:gHom_comp_gsurjec gHom_comp_ginjec)
done

lemma gbij_automorph:"[|Group G; gbijG,G f; gbijG,G g|]  ==> 
                            gbijG,G (g oG f)"
apply (simp add:gbij_comp_bij)
done

lemma l_unit_gHom:"[|Group F; Group G; f ∈ gHom F G|] ==> (IG) oF f = f"
apply (simp add:cmpghom_def) 
apply (rule funcset_eq[of _ "carrier F"],
       simp add:compose_def, simp add:gHom_def)

apply (rule ballI,
       simp add:idmap_def compose_def,
       simp add:gHom_mem[of "F" "G" "f"])
done

lemma r_unit_gHom:"[|Group F; Group G; f ∈ gHom F G |] ==> f oF (IF) = f"
apply (simp add:cmpghom_def)
apply (rule funcset_eq[of _ "carrier F"],
       simp add:compose_def, simp add:gHom_def)
apply (rule ballI,
       simp add:idmap_def compose_def) 
done

section "9. Image"

lemma inv_gHom:"[|Group F; Group G; gbijF,G f|] ==> (Ifn F G f) ∈ gHom G F"
apply (simp add:gHom_def) 
apply (rule conjI)
 (** Ifn F G f ∈ carrier G -> carrier F **) 
apply (simp add:invfun_def restrict_def extensional_def)
apply (rule conjI)
apply (rule inv_func)
 apply (simp add:gbijec_def gsurjec_def gHom_def)
 apply (simp add:gbijec_def gsurjec_def ginjec_def)+

apply (rule ballI)+ apply (erule conjE)+
apply (frule gHom_func[of "F" "G" "f"], assumption+,
       frule inv_func[of "f" "carrier F" "carrier G"], assumption+)
apply (frule_tac b = x in invfun_mem[of "f" "carrier F" "carrier G"],
                 assumption+,
       frule_tac b = y in invfun_mem[of "f" "carrier F" "carrier G"],
                 assumption+)
apply (frule_tac x = "(Ifn F G f) x" and y = "(Ifn F G f) y" in 
                      gHom[of "F" "G" "f"], assumption+)
apply (simp only:invfun_r)
apply (frule sym, thin_tac "f ((Ifn F G f) x ·F (Ifn F G f) y) = x ·G y")
apply (frule_tac a = x and b = y in Group.mult_closed[of "G"], assumption+)
apply (frule_tac b = "x ·G y" in invfun_r[of "f" "carrier F" 
                        "carrier G"], assumption+)
 apply (frule_tac r = "f ((Ifn F G f) (x ·G y))" and s = "x ·G y" and t = "f ((Ifn F G f) x ·F (Ifn F G f) y)" in trans, assumption+)
 apply (thin_tac "f ((Ifn F G f) (x ·G y)) = x ·G y",
        thin_tac "x ·G y = f ((Ifn F G f) x ·F (Ifn F G f) y)")
 apply (frule_tac b = "x ·G y" in invfun_mem[of "f" "carrier F" "carrier G"],
         assumption+,
        frule_tac a = "(Ifn F G f) x" and b = "(Ifn F G f) y" in 
         Group.mult_closed[of "F"], assumption+)
 apply (simp add:inj_on_def) 
done

lemma inv_gbijec_gbijec:"[|Group F; Group G; gbijF,G f|] ==> gbijG,F (Ifn F G f)"
apply (frule inv_gHom [of "F" "G" "f"], assumption+)
apply (simp add:gbijec_def)
apply (rule conjI)
(* gsurjec *)
apply (simp add:gsurjec_def ginjec_def, (erule conjE)+)
apply (frule gHom_func[of "F" "G" "f"], simp add:invfun_surj,
       assumption+, simp add:invfun_surj)
(* ginjec *)
apply (erule conjE, simp add:gsurjec_def ginjec_def, erule conjE,
       frule gHom_func[of "F" "G" "f"], assumption+,
       rule invfun_inj, assumption+)
done

lemma l_inv_gHom:"[|Group F; Group G; gbijF,G f|] ==> (Ifn F G f) oF f = (IF)"
apply (rule funcset_eq[of _ "carrier F"],
       simp add:cmpghom_def, simp add:idmap_def,
       rule ballI)
 apply (simp add:idmap_def cmpghom_def compose_def,
        simp add:gbijec_def gsurjec_def ginjec_def, (erule conjE)+,
        frule gHom_func[of "F" "G" "f"], assumption+)
 apply (rule invfun_l, assumption+)
done
  
lemma img_mult_closed:"[|Group F; Group G; f ∈ gHom F G; u ∈ f `(carrier F);
v ∈ f `(carrier F)|] ==> u ·G v ∈ f `(carrier F)" 
apply (simp add:image_def)
apply ((erule bexE)+, simp) 
apply (subst gHom [of "F" "G" "f", THEN sym], assumption+)
apply (frule_tac a = x and b = xa in Group.mult_closed [of "F"], assumption+)
apply blast
done

lemma img_unit_closed:"[|Group F; Group G; f ∈ gHom F G|] ==>
                                            \<one>G ∈ f `(carrier F)"
apply (cut_tac Group.unit_closed[of "F"],
      frule ghom_unit_unit[THEN sym, of "F" "G" "f"], assumption+,
      simp add:image_def, blast,
      assumption)
done

lemma imgTr7:"[|Group F; Group G; f ∈ gHom F G; u ∈ f `(carrier F)|]
  ==> ρG u  ∈ f `(carrier F)"
apply (simp add:image_def, erule bexE, simp)
 apply (subst ghom_inv_inv[THEN sym, of "F" "G" "f"], assumption+)
 apply (frule_tac a = x in Group.i_closed[of "F"], assumption+)
 apply blast
done

lemma imgTr8:"[|Group F; Group G; f ∈ gHom F G; F » H; u ∈ f ` H; 
                v ∈ f` H |] ==> u ·G v ∈ f ` H" 
apply (simp add:image_def, (erule bexE)+, simp,
       subst gHom [of "F" "G" "f" _, THEN sym], assumption+)
 apply (simp only:Group.sg_subset_elem[of "F"],
         simp only:Group.sg_subset_elem[of "F"])
 apply (frule_tac x = x and y = xa in Group.sg_mult_closed[of "F" "H"],
           assumption+)
 apply blast
done

lemma imgTr9:"[|Group F; Group G; f ∈ gHom F G; F » H; u ∈ f ` H|] ==> 
                    ρG u ∈ f ` H" 
apply (simp add:image_def, erule bexE, simp)
apply (frule_tac h = x in Group.sg_subset_elem[of "F" "H"], assumption+,
       simp add:ghom_inv_inv[THEN sym])
apply (frule_tac x = x in Group.sg_i_closed [of "F" "H"], assumption+,
       blast)
done

lemma imgTr10:"[|Group F; Group G; f ∈ gHom F G; F » H|] ==> \<one>G ∈ f ` H"
apply (frule Group.sg_unit_closed[of "F" "H"], assumption+,
       subst ghom_unit_unit[THEN sym, of "F" "G" "f"], assumption+)
apply (simp add:image_def, blast)
done

lemma imgTr11:"[|Group F; Group G; f ∈ gHom F G; F » H|] ==> G » (f ` H)"
apply (frule gHom_func[of "F" "G" "f"], assumption+,
       frule Group.sg_subset[of "F" "H"], assumption+,
       frule image_sub[of "f" "carrier F" "carrier G" "H"], assumption+)
apply (rule Group.sg_condition[of "G" "f ` H"], assumption+,
       frule imgTr10[of "F" "G" "f" "H"], assumption+,
       rule nonempty, assumption)
apply ((rule allI)+, rule impI, erule conjE,
       frule_tac u = b in imgTr9[of "F" "G" "f" "H"], assumption+,
       frule_tac u = a and v = "ρG b" in imgTr8[of "F" "G" "f"], assumption+)
done

lemma sg_gimg:"[|Group F; Group G; f ∈ gHom F G |]  ==> G » f`(carrier F)"  
apply (frule Group.special_sg_G [of "F"])
apply (simp add:imgTr11)
done

lemma Group_Img:"[|Group F; Group G; f ∈ gHom F G |] ==> Group (ImgF,G f)"
apply (frule sg_gimg [of "F" "G" "f"], assumption+,
       simp add:Gimage_def Group.Group_Gp[of "G"])
done

lemma Img_carrier:"[|Group F; Group G; f ∈ gHom F G |] ==> 
                         carrier (ImgF,G f) = f ` (carrier F)"
by (simp add:Gimage_def Gp_def)

lemma hom_to_Img:"[|Group F; Group G; f ∈ gHom F G|] ==> f ∈ gHom F (ImgF,G f)"
apply (simp add:gHom_def Gimage_def Gp_def,
       (erule conjE)+)
apply (simp add:Pi_def restrict_def)
done

lemma gker_hom_to_img:"[|Group F; Group G; f ∈ gHom F G |] ==>
                               gkerF,(ImgF,G f) f = gkerF,G f" 
by (simp add:gkernel_def Gimage_def,
       frule sg_gimg[of "F" "G" "f"], assumption+,
       simp add:Group.one_Gp_one[of "G" "f ` (carrier F)"])

lemma Pj_im_subg:"[|Group G; G » H; G \<triangleright> K; K ⊆ H|] ==> 
                         Pj G K ` H = carrier ((Gp G H) / K)"
apply (simp add:Qg_def [of "Gp G H" "K"])
apply (rule equalityI, rule subsetI)
 apply (simp add:image_def Pj_def set_rcs_def)
 apply (erule bexE,
        frule_tac h = xa in Group.sg_subset_elem[of "G" "H"], assumption+,
        simp)
apply (simp add:Group.Gp_carrier)
 apply (frule Group.nsg_sg[of "G" "K"], assumption+)
 apply (frule_tac x1 = xa in Group.Gp_rcs[THEN sym, of "G" "K" "H"], 
         assumption+, blast)

apply (rule subsetI)
 apply (simp add:image_def Pj_def)
 apply (simp add:set_rcs_def)
 apply (simp add:Group.Gp_carrier, erule bexE)
  apply (frule Group.nsg_sg[of "G" "K"], assumption+)
 apply (frule_tac x = a in Group.Gp_rcs[of "G" "K" "H"], assumption+,
        simp)
 apply (frule_tac h = a in Group.sg_subset_elem[of "G" "H"], assumption+)  
 apply blast
done

lemma (in Group) subg_Qsubg:"[|G » H; G \<triangleright> K; K ⊆ H|] ==> 
                            (G / K) »  carrier ((Gp G H) / K)"
apply (frule Pj_ghom[of "K"])
apply (frule nsg_sg [of "K"])
apply (frule Group_Qg[of "K"])
apply (cut_tac imgTr11 [of "G" "G / K" "Pj G K" "H"])
apply (cut_tac Pj_im_subg [of "G" "H" "K"])
apply simp apply (rule Group_axioms | assumption)+
done

section "10 incuded homomorphisms"

lemma inducedhomTr:"[|Group F; Group G; f ∈ gHom F G; 
  S ∈ set_rcs F (gkerF,G f); s1 ∈ S; s2 ∈ S |] ==> f s1 = f s2"
apply (simp add:set_rcs_def, erule bexE) 
apply (frule_tac a1 = a in gkernTr11[THEN sym, of "F" "G" "f"], assumption+,
       simp, thin_tac "S = iim F G f {f a}",
       thin_tac "gkerF,G f •F a = iim F G f {f a}")
apply (simp add:iim_def)
done

constdefs
 induced_ghom ::"[('a, 'more) Group_scheme, ('b, 'more1) Group_scheme, 
                ('a  => 'b)] => ('a set  => 'b )"
  "induced_ghom F G f == λX∈ (set_rcs F (gkerF,G f)). f (SOME  x. x ∈ X)"

syntax 
 "@INDUCED_GHOM"::"['a => 'b, ('a, 'm) Group_scheme, ('b, 'm1) Group_scheme]
 =>  ('a set  => 'b )" ("(3_¨_,_)" [82,82,83]82)  

translations
    "f¨F,G " == "induced_ghom F G f"

lemma induced_ghom_someTr:"[|Group F; Group G; f ∈ gHom F G; 
X ∈ set_rcs F (gkerF,G f)|] ==> f (SOME xa. xa ∈ X) ∈ f `(carrier F)"
apply (simp add:set_rcs_def, erule bexE, simp)
apply (frule gkernTr7 [of "F" "G" "f"], assumption+)
apply (rule someI2_ex)
apply (frule_tac a = a in Group.a_in_rcs[of "F" "gkerF,G f"], assumption+) 
apply blast
 apply (frule_tac a = a and x = x in Group.rcs_subset_elem[of "F" "gkerF,G f"],
        assumption+)
 apply (simp add:image_def, blast)
done

lemma induced_ghom_someTr1:"[|Group F; Group G; f ∈ gHom F G; a ∈ carrier F|] ==>
         f (SOME xa. xa ∈ (gkerF,G f) •F a) = f a"
apply (rule someI2_ex)
 apply (frule gkernTr7 [of "F" "G" "f"], assumption+)
 apply (frule Group.a_in_rcs [of "F" "gkerF,G f" "a"], assumption+)
 apply blast
 apply (simp add:gkernTr11 [THEN sym])
 apply (simp add:iim_def)
done

lemma inducedHOMTr0:"[|Group F; Group G; f ∈ gHom F G; a ∈ carrier F|] ==>
     (f¨F,G) ((gkerF,G f) •F a) = f a"
apply (simp add:induced_ghom_def)
apply (frule gkernTr7[of "F" "G" "f"], assumption+)
 apply (simp add:Group.rcs_in_set_rcs[of "F"])
 apply (simp add:induced_ghom_someTr1)
done

lemma inducedHOMTr0_1:"[|Group F; Group G; f ∈ gHom F G|] ==>
                        (f¨F,G) ∈  set_rcs F (gkerF,G f) -> carrier G" 
apply (rule univar_func_test)
 apply (rule ballI, simp add:set_rcs_def, erule bexE, simp)
 apply (subst inducedHOMTr0[of "F" "G" "f"], assumption+)
 apply (simp add:gHom_mem)
done

lemma inducedHOMTr0_2:"[|Group F; Group G; f ∈ gHom F G|] ==>
    (f¨F,G) ∈ set_rcs F (gkerF,G f) -> f ` (carrier F)"
apply (rule univar_func_test, rule ballI)
 apply (simp add:set_rcs_def, erule bexE, simp)
 apply (subst inducedHOMTr0[of "F" "G" "f"], assumption+)
 apply (simp add:image_def, blast)
done

lemma inducedHom:"[|Group F; Group G; f ∈ gHom F G|] ==> 
                         (f¨F,G) ∈ gHom (F/(gkerF,G f)) G"
apply (simp add: gHom_def [of "F/(gkerF,G f)" "G"])
apply (rule conjI)
apply (simp add:induced_ghom_def extensional_def)
 apply (rule allI) apply (rule impI)+ apply (simp add:Qg_def)

apply (rule conjI)
 apply (simp add:Qg_def inducedHOMTr0_1)
apply (rule ballI)+ 
apply (simp add:Qg_def set_rcs_def, (erule bexE)+)
apply simp
 apply (frule gker_normal[of "F" "G" "f"], assumption+)
 apply (simp add:Group.c_top_welldef [THEN sym, of "F" "gkerF,G f"])
 apply (frule_tac a = a and b = aa in Group.mult_closed[of "F"], assumption+)
 apply (simp add:inducedHOMTr0)
apply (simp add:gHom)
done

lemma induced_ghom_ginjec: "[|Group F; Group G; f ∈ gHom F G |] ==> 
            ginj(F/(gkerF,G f)),G (f¨F,G)"
apply (simp add:ginjec_def)
apply (simp add:inducedHom)
apply (simp add:inj_on_def)
apply (rule ballI)+
apply (simp add:Qg_def)
apply (simp add:set_rcs_def)
apply ((erule bexE)+, rule impI, simp)
apply (simp add:inducedHOMTr0)
apply (simp add: gkernTr11[THEN sym])
done

lemma inducedhomgsurjec:"[|Group F; Group G; gsurjF,G f|] ==>
                                gsurj(F/(gkerF,G f)),G (f¨F,G)"
apply (simp add:gsurjec_def)
apply (simp add:inducedHom)
 apply (erule conjE)
 apply (frule gHom_func[of "F" "G" "f"], assumption+)
 apply (rule surj_to_test)
 apply (simp add:Qg_def)
 apply (frule inducedHOMTr0_2[of "F" "G" "f"], assumption+)
 apply (simp add:surj_to_def[of "f" "carrier F" "carrier G"])

apply (rule ballI)
 apply (simp add:surj_to_def[of "f" "carrier F" "carrier G"],
        frule sym, thin_tac "f ` carrier F = carrier G", simp,
        thin_tac "f ∈ carrier F -> f ` carrier F",
        thin_tac "carrier G = f ` carrier F")
 apply (simp add:image_def, erule bexE, simp,
        thin_tac "b = f x")
 apply (simp add:Qg_def)
 apply (frule_tac a = x in inducedHOMTr0[of "F" "G" "f"], assumption+)
 apply (frule gkernTr7[of "F" "G" "f"], assumption+)
 apply (frule_tac a = x in Group.rcs_in_set_rcs[of "F" "gkerF,G f"],
                              assumption+)
 apply blast
done

lemma homomtr: "[|Group F; Group G; f ∈ gHom F G|] ==> 
                  (f¨F,G) ∈ gHom (F / (gkerF,G f)) (ImgF,G f)"
apply (simp add:gHom_def [of "F / gkerF,G f" _])
apply (rule conjI)
apply (simp add:induced_ghom_def extensional_def)
 apply (rule allI, (rule impI)+, simp add:Qg_def)
apply (rule conjI)
 apply (rule univar_func_test, rule ballI)
 apply (simp add:Gimage_def Qg_def Gp_def, simp add:set_rcs_def, erule bexE)
 apply simp
 apply (simp add:inducedHOMTr0)

apply (rule ballI)+
 apply (simp add:Qg_def set_rcs_def, (erule bexE)+, simp)
 apply (frule gker_normal[of "F" "G" "f"], assumption+)
 apply (simp add:Group.c_top_welldef[THEN sym, of "F" "gkerF,G f"])
 apply (frule_tac a = a and b = aa in Group.mult_closed[of "F"], assumption+)
 apply (simp add:inducedHOMTr0)
 apply (simp add:Gimage_def)
 apply (subst Group.Gp_mult_induced[of "G" "f ` carrier F"], assumption+)
  apply (simp add:sg_gimg)
  apply (simp add:image_def, blast)
  apply (simp add:image_def, blast)
apply (simp add:gHom)
done

lemma homom2img: "[|Group F; Group G; f ∈ gHom F G |] ==> 
   (f¨F,(ImgF,G f)) ∈ gHom (F / (gkerF,G f)) (ImgF,G f)"
apply (frule hom_to_Img[of "F" "G" "f"], assumption+)
apply (frule inducedHom[of "F" "ImgF,G f" "f"])
 apply (simp add:Group_Img) apply assumption
 apply (simp add:gker_hom_to_img[of "F" "G" "f"])
done

lemma homom2img1:"[|Group F; Group G; f ∈ gHom F G; X ∈ set_rcs F (gkerF,G f)|]
 ==> (f¨F,(ImgF,G f)) X = (f¨F,G) X"
apply (frule gker_hom_to_img [of "F" "G" "f"], assumption+)
apply (simp add:induced_ghom_def)
done

subsection "11 Homomorphism therems"

constdefs
  iota :: "('a, 'm) Group_scheme => ('a => 'a)"  
(* should be used exclusively as an inclusion map *)
          ("(ι_)" [1000]999) 
   "ιF == λx ∈ (carrier F). x"

lemma iotahomTr0:"[|Group G; G » H; h ∈ H |] ==> (ι(Gp G H)) h = h"
apply (simp add:iota_def)
apply (simp add:Gp_def)
done
     
lemma iotahom:"[|Group G; G » H; G \<triangleright> N|] ==> 
                ι(Gp G H) ∈ gHom (Gp G H) (Gp G (H \<struct>G N))" 
apply (simp add:gHom_def)
apply (rule conjI)
 apply (simp add:iota_def extensional_def)
apply (rule conjI)
apply (simp add:Pi_def restrict_def iota_def)
 apply (rule allI, rule impI)
 apply (simp add:Gp_def)
 apply (frule Group.nsg_sg[of "G" "N"], assumption+) 
 apply (frule Group.l_sub_smult[of "G" "H" "N"], assumption+, 
                                                   simp add: subsetD)
apply (rule ballI)+
 apply (simp add:iota_def, simp add:Group.Gp_carrier)
 apply (frule Group.smult_sg_nsg[of "G" "H" "N"], assumption+,
        frule Group.l_sub_smult[of "G" "H" "N"], assumption+,
        simp add:Group.nsg_sg,
        frule_tac c = x in subsetD[of "H" "H \<struct>G N"], assumption+,
        frule_tac c = y in subsetD[of "H" "H \<struct>G N"], assumption+)
 apply (simp add:Group.Gp_mult_induced[of "G"])
  apply (simp add:Group.sg_mult_closed)
done

lemma iotaTr0: "[|Group G; G » H; G \<triangleright> N|] ==> 
               ginj(Gp G H),(Gp G (H \<struct>G N))(Gp G H))"
apply (simp add:ginjec_def)
apply (simp add:iotahom)
apply (simp add:inj_on_def iota_def Gp_def)
done

theorem homomthm1:"[|Group F; Group G; f ∈ gHom F G |] ==>
     gbij(F/ (gkernel F G f)), (Gimage F G f) (f¨F, (Gimage F G f))"
apply (frule homom2img [of "F" "G" "f"], assumption+)
apply (simp add:gbijec_def)
apply (frule hom_to_Img [of "F" "G" "f"], assumption+)
apply (frule Group_coim[of "F" "G" "f"], assumption+,
       frule gHom_func[of "F / gkerF,G f" "ImgF,G f" 
        "f¨F,ImgF,G f"], simp add:Group_Img, assumption)
apply (rule conjI)
 (** gsurjec **)  
 apply (simp add:gsurjec_def,
        rule surj_to_test[of "f¨F,ImgF,G f" 
        "carrier (F / gkerF,G f)" "carrier (ImgF,G f)"], assumption+,
        rule ballI)
 apply (thin_tac "f¨F,ImgF,G f ∈ gHom (F / gkerF,G f)
                        (ImgF,G f)") 
 apply (simp add:Img_carrier, simp add:image_def, erule bexE, simp,
        simp add:Qg_def,
        frule gkernTr7[of "F" "G" "f"], assumption+)
 apply (frule_tac a = x in Group.rcs_in_set_rcs[of "F" "gkerF,G f"],
               assumption+)
 apply (simp add:homom2img1) 
 apply (frule_tac a = x in inducedHOMTr0[of "F" "G" "f"], assumption+)
 apply blast  (** gsurjec done **)
(** ginjec **)
apply (frule induced_ghom_ginjec[of "F" "G" "f"], assumption+)
 apply (simp add:ginjec_def) 
 apply (frule conjunct2)
  apply (thin_tac "f¨F,G ∈ gHom (F / gkerF,G f) G ∧
     inj_on (f¨F,G) (carrier (F / gkerF,G f))")
 apply (simp add:inj_on_def)
apply ((rule ballI)+, rule impI) 
apply (simp add:Qg_def, fold Qg_def)
apply (simp add:homom2img1) 
done

lemma isomTr0 [simp]:"Group F ==> F ≅ F"
by (frule IdTr2 [of "F"], simp add:isomorphic_def,
       blast)

lemma isomTr1:"[|Group F; Group G; F ≅  G |] ==> G ≅ F"
apply (simp add:isomorphic_def, erule exE)
apply (frule_tac f = f in inv_gbijec_gbijec[of "F" "G"], assumption+)
apply blast
done
 
lemma isomTr2:"[|Group F; Group G; Group H; F ≅ G; G ≅ H |] ==> F ≅ H"
apply (simp add:isomorphic_def)
apply (erule exE)+
apply (simp add:gbijec_def)
 apply (erule conjE)+
apply (frule gHom_comp_gsurjec [of "F" "G" "H" _ _], assumption+)
apply (frule gHom_comp_ginjec [of "F" "G" "H" _ _], assumption+)
apply blast
done
 
lemma gisom1: "[|Group F; Group G; f ∈ gHom F G |] ==>
     (F/ (gkerF,G f)) ≅ (ImgF,G f)"
apply (simp add:isomorphic_def)
apply (frule homomthm1 [of "F" "G" "f"], assumption+)
apply blast
done

lemma homomth2Tr0: "[|Group F; Group G; f ∈ gHom F G; G \<triangleright> N|] ==>   
                           F \<triangleright> (iim F G f N)"
apply (frule Group.cond_nsg[of "F" "iim F G f N"],
       frule Group.nsg_sg[of "G" "N"], assumption+,
       simp add:ghomTr4[of "F" "G" "f" "N"])
apply ((rule ballI)+,
       simp add:iim_def, erule conjE,
       frule_tac a = a in Group.i_closed[of "F"], assumption+,
       frule_tac a = a and b = h in Group.mult_closed[of "F"], assumption+)
 apply (simp add:gHom ghom_inv_inv Group.mult_closed)
apply (frule_tac x = a in gHom_mem[of "F" "G" "f"], assumption+,
       simp add:Group.nsgPr1_1,
       assumption)
done

lemma kern_comp_gHom:"[|Group F; Group G; gsurjF,G f; G \<triangleright> N|] ==>
                  gkerF, (G/N) ((Pj G N) oF f) = iim F G f N"
apply (simp add:gkernel_def iim_def)
apply (simp add:Group.Qg_one[of "G" "N"] cmpghom_def compose_def)
apply (rule equalityI)
 apply (rule subsetI, simp, erule conjE, simp)
 apply (simp add:gsurjec_def, frule conjunct1, fold gsurjec_def)
 apply (frule_tac x = x in gHom_mem[of "F" "G" "f"], assumption+)
 apply (simp add:Group.Pj_mem[of "G" "N"])
 apply (frule Group.nsg_sg[of "G" "N"], assumption+)
 apply (frule_tac a = "f x" in Group.a_in_rcs[of "G" "N"], assumption+)
 apply simp
apply (rule subsetI)
 apply (simp, erule conjE)
 apply (frule Group.nsg_sg[of "G" "N"], assumption,
        frule_tac h = "f x" in Group.sg_subset_elem[of "G" "N"], assumption+)
 apply (simp add:Group.Pj_mem[of "G" "N"])
 apply (simp add:Group.rcs_fixed2[of "G" "N"])
done

lemma QgrpUnit_1:"[|Group G; Ugp E; G \<triangleright> H; (G / H) ≅ E |] ==> carrier G = H"
apply (simp add:isomorphic_def, erule exE)
apply (frule Group.Group_Qg[of "G" "H"], assumption+,
       simp add:gbijec_def, erule conjE)
apply (frule_tac f = f in gkern2[of "G / H" "E"],
       simp add:Ugp_def, simp add:gsurjec_def, assumption,
       simp add:gsurjec_def, frule conjunct1, fold gsurjec_def,
       frule_tac f = f in gkern1[of "G/H" "E"], assumption+)
 apply (simp, thin_tac "gker(G / H),E f = {\<one>G / H}",
        thin_tac "gsurj(G / H),E f", thin_tac "ginj(G / H),E f",
        thin_tac "Group (G / H)", thin_tac "f ∈ gHom (G / H) E",
        simp add:Group.Qg_carrier)
apply (rule contrapos_pp, simp+,
       frule Group.nsg_sg[of "G" "H"], assumption+,
       frule Group.sg_subset[of "G" "H"], assumption+,
       frule sets_not_eq[of "carrier G" "H"], assumption, erule bexE,
       frule_tac a = a in Group.rcs_in_set_rcs[of "G" "H"], assumption+,
       simp)
 apply (thin_tac "set_rcs G H = {\<one>G / H}", simp add:Qg_def,
        frule_tac a = a in Group.a_in_rcs[of "G" "H"], assumption+,
        simp)
done
    
lemma QgrpUnit_2:"[|Group G; Ugp E; G \<triangleright> H; carrier G = H|] ==> (G/H) ≅ E"
apply (frule Group.Group_Qg [of "G" "H"], assumption+)
apply (simp add:Group.Qg_unit_group[THEN sym, of "G" "H"])
apply (simp add:Ugp_def)
apply (frule Group.Qg_carrier[of "G" "H"], simp) 
apply (thin_tac "set_rcs G H = {H}")
 apply (frule Group.Qg_one[of "G" "H"], assumption+, erule conjE) 
 apply (rule Ugps_isomorphic[of "G / H" "E"])
 apply (simp add:Ugp_def)+
done

lemma QgrpUnit_3:"[|Group G; Ugp E; G » H; G » H1; (Gp G H) \<triangleright> H1;
                    ((Gp G H) / H1) ≅ E |] ==> H = H1"
apply (frule Group.Group_Gp[of "G" "H"], assumption+)
apply (frule QgrpUnit_1 [of "Gp G H" "E" "H1"], assumption+)
apply (simp add:Group.Gp_carrier)
done

lemma QgrpUnit_4:"[|Group G; Ugp E; G » H; G » H1; (Gp G H) \<triangleright> H1;
¬ ((Gp G H) / H1) ≅ E |] ==> H ≠ H1"
apply (frule Group.Group_Gp[of "G" "H"], assumption+)
apply (rule contrapos_pp, simp) apply simp
apply (frule QgrpUnit_2 [of "Gp G H1" "E" "H1"], assumption+)
apply (simp add:Group.Gp_carrier)
apply simp
done

constdefs
  Qmp :: "[('a, 'm) Group_scheme, 'a set, 'a set] => ('a set => 'a set)"
   "Qmp G H N == λX∈ set_rcs G H. {z. ∃ x ∈ X. ∃ y ∈ N. (y ·G x = z)}"
  
syntax
   "@QP" :: "['a Group, 'a set, 'a set] => ('a set => 'a set)"
               ("(3Qm_ _,_)" [82,82,83]82)
translations
   "QmG H,N" == "Qmp G H N"

 (* "[| Group G; G \<triangleright> H; G \<triangleright> N; H ⊆ N |] ==>
           QmpG H,N  ∈ gHom (G / H) (G / N)"  *)

  (* show Qmp G H N is a welldefined map from G/H to G/N. step1 *)

lemma (in Group) QmpTr0:"[|G » H; G » N; H ⊆ N ; a ∈ carrier G|] ==>
                                          Qmp G H N (H • a) = (N • a)"
apply (frule_tac a = a in rcs_in_set_rcs[of "H"], assumption,
       simp add:Qmp_def)
apply (rule equalityI)
 apply (rule subsetI, simp, (erule bexE)+,
        thin_tac "H • a ∈ set_rcs G H",
        simp add:rcs_def, erule bexE)
 apply (frule sym, thin_tac "y · xa = x", frule sym, thin_tac "h · a = xa",
        simp,
        frule_tac h = y in sg_subset_elem[of "N"], assumption+,
        frule_tac h = h in sg_subset_elem[of "H"], assumption+,
        frule_tac c = h in subsetD[of "H" "N"], assumption+,
        frule_tac x = y and y = h in sg_mult_closed[of "N"], assumption+,
        subst tassoc[THEN sym], assumption+, blast)

apply (rule subsetI,
       thin_tac "H • a ∈ set_rcs G H",
       simp add:rcs_def, erule bexE,
       frule sg_unit_closed[of "H"],
       frule l_unit[of "a"], blast) 
done

  (* show Qmp G H N is a welldefined map from G/H to G/N. step2 *)
lemma (in Group) QmpTr1:"[|G » H; G » N; H ⊆ N; a ∈ carrier G; b ∈ carrier G; 
                      H • a = H • b|] ==> N • a = N • b"
apply (simp add:rcs_eq[THEN sym, of "H" "a" "b"],
       frule subsetD[of "H" "N" "b · ρ a"], assumption+,
       simp add:rcs_eq[of "N" "a" "b"])
done
   
lemma (in Group) QmpTr2:"[|G » H; G » N; H ⊆ N ; X ∈ carrier (G/H)|]
                        ==> (Qmp G H N) X ∈ carrier (G/N)" 
by (simp add:Qg_carrier[of "H"] set_rcs_def, erule bexE, simp add: QmpTr0,
       simp add:Qg_carrier rcs_in_set_rcs)

lemma (in Group) QmpTr2_1:"[|G » H; G » N; H ⊆ N |] ==> 
                           Qmp G H N ∈ carrier (G/H) -> carrier (G/N)"
by (rule univar_func_test, rule ballI, simp add:QmpTr2 [of "H" "N"])

lemma (in Group) QmpTr3:"[|G \<triangleright> H; G \<triangleright> N; H ⊆ N; X ∈ carrier (G/H); 
       Y ∈ carrier (G/H)|] ==> 
     (Qmp G H N) (c_top G H X Y) = c_top G N ((Qmp G H N) X) ((Qmp G H N) Y)" 
apply (frule nsg_sg[of "H"], frule nsg_sg[of "N"])
apply (simp add:Qg_carrier)
apply (simp add:set_rcs_def, (erule bexE)+, simp)
apply (subst c_top_welldef [THEN sym], assumption+)
apply (frule_tac  a = a and b = aa in mult_closed, assumption+)
apply (simp add:QmpTr0)+
apply (subst c_top_welldef [THEN sym], assumption+)
apply simp
done
 
lemma (in Group) Gp_s_mult_nsg:"[|G \<triangleright> H; G \<triangleright> N; H ⊆ N; a ∈ N |] ==>
                                 H •(Gp G N) a =  H • a"
by (frule nsg_sg[of "H"], frule nsg_sg[of "N"], 
                          rule Gp_rcs[of "H" "N" "a"], assumption+)

lemma (in Group) QmpTr5:"[|G \<triangleright> H; G \<triangleright> N; H ⊆ N; X ∈ carrier (G/H); 
      Y ∈ carrier (G/H) |] ==> (Qmp G H N) ( X ·(G / H) Y) =
                              ((Qmp G H N) X) ·(G / N) ((Qmp G H N) Y)"
by (frule nsg_sg[of "H"], frule nsg_sg[of "N"],
       (subst Qg_def)+, simp,
       simp add:QmpTr3 [of "H" "N" "X" "Y"])

lemma (in Group) QmpTr:"[|G \<triangleright> H; G \<triangleright> N; H ⊆ N |] ==>
                            (QmG H,N) ∈ gHom (G / H) (G / N)"
apply (simp add:gHom_def)
apply (rule conjI)  
 apply (simp add:Qmp_def extensional_def)
apply (rule allI, (rule impI)+, simp add:Qg_def) 
apply (rule conjI) 
 apply (rule QmpTr2_1[of "H" "N"])
 apply (simp add:nsg_sg)+
apply (rule ballI)+ 
 apply (simp add:QmpTr5)
done
     
lemma (in Group) Qmpgsurjec:"[|G \<triangleright> H; G \<triangleright> N; H ⊆ N |] ==> 
                                      gsurj(G / H),(G / N) (QmG H,N)"
apply (frule nsg_sg[of "H"], frule nsg_sg[of "N"])  
apply (frule QmpTr [of "H" "N"], assumption+) 
apply (simp add:gsurjec_def)
apply (rule surj_to_test)
 apply (simp add:QmpTr2_1)
apply (rule ballI)
 apply (simp add:Qg_carrier,
        simp add:set_rcs_def[of "G" "N"], erule bexE,
        frule_tac a = a in QmpTr0[of "H" "N"], assumption+, simp)
 apply (frule_tac a = a in rcs_in_set_rcs[of "H"], assumption+,
        blast)
done
(*
constdefs
  S_d::"[('a, 'm) Group_scheme, 'a set, 'a set] => 'a set set"
   "S_d G N H == carrier ((Gp G N) / H)"
  
syntax
   "@S_D":: "['a set, ('a, 'm) Group_scheme, 'a set] => 'a set set"
               ("(3_ '/'s_ _)" [82,82,83]82)
translations
   "N /sG H" == "S_d G N H" *)

lemma (in Group) gkerQmp:"[|G \<triangleright> H; G \<triangleright> N; H ⊆ N |] ==>
               gker(G / H),(G / N) (QmG H,N) = carrier ((Gp G N)/ H)"   
apply (frule nsg_sg[of "H"], frule nsg_sg[of "N"])
apply (simp add:gkernel_def)
apply (rule equalityI)
 apply (rule subsetI,
        simp add:Qg_carrier set_rcs_def, erule conjE, erule bexE, simp,
        simp add:Qg_one) 
 apply (simp add:QmpTr0,
        frule_tac a = a in a_in_rcs[of "N"], assumption+, simp,
        frule Group_Gp[of "N"])
 apply (simp add:Group.Qg_carrier, simp add:set_rcs_def, simp add:Gp_carrier,
        simp add:Gp_rcs, blast)

apply (rule subsetI)
 apply (frule Group_Gp[of "N"],
        simp add:Group.Qg_carrier Qg_one set_rcs_def, erule bexE,
        simp add:Qg_carrier, thin_tac "x = H •\<natural>N a")
 apply (simp add:Gp_carrier, simp add:Gp_rcs,
        frule_tac h = a in sg_subset_elem[of "N"], assumption,
        simp add:rcs_in_set_rcs, simp add:QmpTr0,
        simp add:rcs_fixed2[of "N"])
done

theorem (in Group) homom2:"[|G \<triangleright> H; G \<triangleright> N; H ⊆ N|] ==>
          gbij((G/H)/(carrier ((Gp G N)/H))),(G/N) ((QmG H,N(G/H),(G/N))"
apply (frule QmpTr [of "H" "N"], assumption+)
apply (simp add:gbijec_def)
apply (rule conjI)
apply (frule Group_Qg[of "H"], frule Group_Qg[of "N"])
apply (frule inducedHom [of "G/H" "G/N" " Qmp G H N"], assumption+)
(** show  surj_to **)
apply (frule Qmpgsurjec [of "H" "N"], assumption+)
apply (frule inducedhomgsurjec [of "G/H" "G/N" "QmG H,N"], assumption+)
apply (simp add:gkerQmp [of "H" "N"])

(** show ginj **)
apply (frule QmpTr [of "H" "N"], assumption+)
apply (frule Group_Qg[of "H"], frule Group_Qg[of "N"])
apply (frule induced_ghom_ginjec [of "G/H" "G/N" "Qmp G H N"], assumption+)
apply (simp add:gkerQmp [of "H" "N"])
done

section "12. Isomorphims"

theorem (in Group) isom2:"[|G \<triangleright> H; G \<triangleright> N; H ⊆ N|] ==>
               ((G/H)/(carrier ((Gp G N)/H))) ≅  (G/N)"
apply (frule homom2 [of "H" "N"], assumption+)
 apply (simp add:isomorphic_def)
 apply blast
done

theorem homom3:"[| Group F; Group G; G \<triangleright> N; gsurjF,G f; 
         N1 = (iim F G f) N |] ==> (F / N1) ≅ (G / N)" 
apply (frule Group.Pj_gsurjec [of "G" "N"], assumption+)
apply (frule Group.Group_Qg[of "G" "N"], assumption)
apply (frule gHom_comp_gsurjec [of "F" "G" "G / N" "f" "Pj G N"], assumption+)
apply (frule inducedhomgsurjec [of "F" "G / N" "(Pj G N) oF f"], assumption+)
apply (frule induced_ghom_ginjec [of "F" "G / N" "(Pj G N) oF f"], assumption+) 
 apply (simp add:gsurjec_def [of "F" "G / N" "(Pj G N) oF f"])
 apply (simp add:kern_comp_gHom[of "F" "G" "f" "N"])
 apply (frule sym, thin_tac "N1 = iim F G f N", simp)
apply (simp add:isomorphic_def gbijec_def, blast) 
done

lemma (in Group) homom3Tr1:"[|G » H; G \<triangleright> N|] ==> H ∩ N =  
gker(Gp G H),((Gp G (H \<struct>G N))/N) 
               ((Pj (Gp G (H \<struct>G N)) N) o(Gp G H)(Gp G H)))"  
apply (simp add:gkernel_def, frule nsg_sg,
       simp add:Gp_carrier[of "H"],
       frule smult_sg_nsg, assumption+,
       frule Gp_smult_nsg[of "H" "N"], assumption,
       frule Group_Gp [of "H \<struct>G N"])
 apply (simp add:Group.Qg_one[of "Gp G (H \<struct>G N)" "N"],
        simp add:iota_def Gp_carrier, simp add:cmpghom_def compose_def,
        simp add:Gp_carrier) 
apply (rule equalityI)
apply (rule subsetI, simp, erule conjE)
 apply (frule_tac h = x in sg_subset_elem[of "H"], assumption+,
        subst Group.Pj_mem, assumption+,
        simp add:Gp_carrier,
        frule l_sub_smult[of "H" "N"], assumption+,
        rule_tac c = x in subsetD[of "H" "H \<struct>G N"], assumption+)
  apply (frule r_sub_smult[of "H" "N"], assumption+,
         frule_tac c = x in subsetD[of "N" "H \<struct>G N"], assumption+,
         simp add:Gp_rcs[of "N" "H \<struct>G N"])
  apply (simp add:rcs_fixed2)

apply (rule subsetI, simp, erule conjE, simp)
 apply (frule_tac h = x in sg_subset_elem[of "H"], assumption+)
 apply (frule l_sub_smult[of "H" "N"], assumption+,
        frule r_sub_smult[of "H" "N"], assumption+)
 apply (frule_tac x = x in Group.Pj_mem[of "Gp G (H \<struct>G N)" "N"], assumption+)
       apply (simp add:Gp_carrier)
       apply (
        rule_tac c = x in subsetD[of "H" "H \<struct>G N"], assumption+)
 apply (frule_tac c = x in subsetD[of "H" "H \<struct>G N"], assumption+)
 apply (simp only:Group.Gp_rcs)
 apply (simp only:Gp_rcs[of "N" "H \<struct>G N"])
 apply (frule_tac a = x in a_in_rcs[of "N"], assumption+, simp)
done     

subsection " An automorphism groups"

constdefs (structure G)
   automg :: "_  => 
      (| carrier :: ('a => 'a) set, top :: ['a => 'a,'a => 'a] => ('a => 'a),
        iop :: ('a => 'a) => ('a => 'a), one :: ('a => 'a)|)),"
 
      "automg G  ==  (| carrier = {f. gbijG,G f}, 
        top = λg∈{f. gbijG,G f}. λf∈{f. gbijG,G f}. ( g oG f), 
        iop = λf∈{f. gbijG,G f}. (Ifn G G f), one = IG |))," 

lemma automgroupTr1:"[|Group G; gbijG,G f; gbijG,G g; gbijG,G h|] ==>
                        (h oG g) oG f =  h oG (g oG f)" 
apply (simp add:cmpghom_def, 
       unfold gbijec_def)
 apply (frule conjunct1, rotate_tac 2, frule conjunct1,
        rotate_tac 1, frule conjunct1, fold gbijec_def)
 apply (simp add:gsurjec_def, (erule conjE)+,
        frule gHom_func[of "G" "G" "f"], assumption+,
        frule gHom_func[of "G" "G" "g"], assumption+,
        frule gHom_func[of "G" "G" "h"], assumption+)
apply (simp add:compose_assoc)
done

lemma automgroup:"Group G  ==> Group (automg G)"
apply (unfold Group_def [of "automg G"])
apply(auto simp: automg_def Pi_def gbij_comp_bij automgroupTr1 IdTr2 Id_l_unit l_inv_gHom inv_gbijec_gbijec)
done

subsection "complete system of representatives"

constdefs (structure G)
   
 gcsrp::"_ => 'a set => 'a set => bool"
 "gcsrp G H S == ∃f. (bij_to f (set_rcs G H) S)"
(** G_csrp is defined iff H is a nsg **)

 gcsrp_map::"_ => 'a set => 'a set => 'a"
 "gcsrp_map G H == λX∈(set_rcs G H). SOME x. x ∈ X"

lemma (in Group) gcsrp_func:"G » H ==> gcsrp_map G H ∈ set_rcs G H -> UNIV"  
apply (rule univar_func_test)
apply (rule ballI)
apply (simp add:set_rcs_def)
done

lemma (in Group) gcsrp_func1:"G » H ==> 
       gcsrp_map G H ∈ set_rcs G H -> (gcsrp_map G H) ` (set_rcs G H)"  
apply (rule univar_func_test)
apply (rule ballI)
apply (simp add:set_rcs_def)
done

lemma (in Group) gcsrp_map_bij:"G » H ==>
         bij_to (gcsrp_map G H) (set_rcs G H) ((gcsrp_map G H) `(set_rcs G H))"
apply (simp add:bij_to_def, rule conjI)
 apply (rule surj_to_test)
 apply (rule univar_func_test, rule ballI)
 apply (simp add:image_def, blast)
apply (rule ballI, simp add:image_def, erule bexE, simp, blast)

apply (simp add:inj_on_def)
 apply ((rule ballI)+, rule impI)
 apply (simp add:gcsrp_map_def)
 apply (frule_tac X = x in rcs_nonempty, assumption+,
        frule_tac X = y in rcs_nonempty, assumption+)
 apply (frule_tac A = x in nonempty_some, 
        frule_tac A = y in nonempty_some, simp)
apply (rule_tac X = x and Y = y in rcs_meet[of "H"], assumption+)
apply blast
done

lemma (in Group) image_gcsrp:"G » H ==> 
                   gcsrp G H ((gcsrp_map G H) `(set_rcs G H))"
apply (simp add:gcsrp_def)
apply (frule gcsrp_map_bij[of "H"], blast)
done

lemma (in Group) gcsrp_exists:"G » H ==> ∃S. gcsrp G H S"
by (frule image_gcsrp[of "H"], blast)

constdefs (structure G)

 gcsrp_top::"[_ , 'a set] =>  'a =>  'a  => 'a"
 "gcsrp_top G H == λx ∈ ((gcsrp_map G H) `(set_rcs G H)).
                     λy ∈ ((gcsrp_map G H) `(set_rcs G H)).
 gcsrp_map G H 
  (c_top G H  
  ((invfun (set_rcs G H) ((gcsrp_map G H) `(set_rcs G H)) (gcsrp_map G H)) x) 
  ((invfun (set_rcs G H) ((gcsrp_map G H) `(set_rcs G H)) (gcsrp_map G H)) y))"

 gcsrp_iop::"[_ , 'a set] => 'a => 'a"
 "gcsrp_iop G H == λx ∈ ((gcsrp_map G H) `(set_rcs G H)).
  gcsrp_map G H 
  (c_iop G H
  ((invfun (set_rcs G H) ((gcsrp_map G H) `(set_rcs G H)) (gcsrp_map G H)) x))"

  gcsrp_one::"[_ , 'a set] => 'a" 
  "gcsrp_one G H == gcsrp_map G H H"

 Gcsrp::"_ => 'a set => 'a Group"
 "Gcsrp G N == (|carrier = (gcsrp_map G N) `(set_rcs G N),
                top = gcsrp_top G N, iop = gcsrp_iop G N, one = gcsrp_one G N|)),"

lemma (in Group) gcsrp_top_closed:"[|Group G; G \<triangleright> N;
  a ∈ ((gcsrp_map G N) `(set_rcs G N)); b ∈ ((gcsrp_map G N) `(set_rcs G N))|]
 ==> gcsrp_top G N a b ∈ (gcsrp_map G N) `(set_rcs G N)"
apply (frule nsg_sg[of "N"],
       frule gcsrp_func1[of "N"],
       frule gcsrp_map_bij[of "N"]) 
apply (frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N"
          "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+,
       frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N"
          "(gcsrp_map G N) ` (set_rcs G N)" "b"], assumption+)
apply (frule Qg_top_closed[of "N" "invfun (set_rcs G N) 
             (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) a"
       "invfun (set_rcs G N) 
             (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) b"], assumption+)
apply (simp add:gcsrp_top_def)
done

lemma (in Group) gcsrp_tassoc:"[|Group G; G \<triangleright> N;
       a ∈ ((gcsrp_map G N) `(set_rcs G N)); 
        b ∈ ((gcsrp_map G N) `(set_rcs G N));
         c ∈ ((gcsrp_map G N) `(set_rcs G N))|] ==>
      (gcsrp_top G N (gcsrp_top G N a b) c) =
               (gcsrp_top G N a (gcsrp_top G N b c))"
apply (frule nsg_sg[of "N"],
       frule gcsrp_func1[of "N"],
       frule gcsrp_map_bij[of "N"]) 
apply (frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N"
          "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+,
       frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N"
          "(gcsrp_map G N) ` (set_rcs G N)" "b"], assumption+,
       frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N"
          "(gcsrp_map G N) ` (set_rcs G N)" "c"], assumption+)
apply (frule Qg_top_closed[of "N" "invfun (set_rcs G N) 
             (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) a"
       "invfun (set_rcs G N) 
             (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) b"], assumption+,
       frule Qg_top_closed[of "N" "invfun (set_rcs G N) 
             (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) b"
       "invfun (set_rcs G N) 
             (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) c"], assumption+) 
apply (simp add:gcsrp_top_def)
 apply (simp add:invfun_l1)
 apply (simp add:Qg_tassoc[of "N"])
done

lemma (in Group) gcsrp_l_one:"[|Group G; G \<triangleright> N;
       a ∈ ((gcsrp_map G N) `(set_rcs G N))|] ==>
             (gcsrp_top G N (gcsrp_one G N) a) = a"
apply (frule nsg_sg[of "N"],
       frule gcsrp_func1[of "N"],
       frule gcsrp_map_bij[of "N"],
       frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N"
          "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+)
apply (simp add:gcsrp_top_def gcsrp_one_def)
 apply (frule Qg_unit_closed[of "N"])
 apply (simp add:funcset_mem)
 apply (simp add:invfun_l1)
 apply (simp add:Qg_unit)
 apply (simp add:invfun_r1)
done

lemma (in Group) gcsrp_l_i:"[|G \<triangleright> N; a ∈ ((gcsrp_map G N) `(set_rcs G N))|] ==>
       gcsrp_top G N (gcsrp_iop G N a) a = gcsrp_one G N"
apply (frule nsg_sg[of "N"],
       frule gcsrp_func1[of "N"],
       frule gcsrp_map_bij[of "N"],
       frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N"
          "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+)
apply (frule Qg_iop_closed[of "N" "invfun (set_rcs G N) (gcsrp_map G N ` 
                       set_rcs G N) (gcsrp_map G N) a"], assumption+) 
apply (simp add:gcsrp_top_def gcsrp_one_def gcsrp_iop_def)
 apply (simp add:invfun_l1 Qg_i)
done

lemma (in Group) gcsrp_i_closed:"[|G \<triangleright> N; a ∈ ((gcsrp_map G N) `(set_rcs G N))|]
      ==> gcsrp_iop G N a ∈ ((gcsrp_map G N) `(set_rcs G N))"
apply (frule nsg_sg[of "N"],
       frule gcsrp_func1[of "N"],
       frule gcsrp_map_bij[of "N"],
       frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N"
          "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+)
apply (frule Qg_iop_closed[of "N" "invfun (set_rcs G N) (gcsrp_map G N ` 
                       set_rcs G N) (gcsrp_map G N) a"], assumption+)  
apply (simp add:gcsrp_iop_def)
done

lemma (in Group) Group_Gcsrp:"G \<triangleright> N ==> Group (Gcsrp G N)"
apply (simp add:Group_def)
 apply (rule conjI)
 apply (rule univar_func_test, rule ballI)
 apply (simp add:Gcsrp_def)
 apply (rule univar_func_test, rule ballI)
 apply (rule_tac a = x and b = xa in gcsrp_top_closed[of "N"], rule Group_axioms, assumption+)
apply (rule conjI)
 apply (rule allI, rule impI)+
 apply (simp add:Gcsrp_def)
 apply (rule_tac a = a and b = b and c = c in gcsrp_tassoc[of "N"], rule Group_axioms, assumption+)
apply (rule conjI)
 apply (rule univar_func_test, rule ballI)
 apply (simp add:Gcsrp_def, rule gcsrp_i_closed[of "N"], assumption+)
apply (rule conjI)
 apply (rule allI, rule impI)
 apply (simp add:Gcsrp_def,
        rule  gcsrp_l_i[of "N"], assumption+)
apply (rule conjI)
 apply (frule Qg_unit_closed[of "N"],
        simp add:Gcsrp_def gcsrp_one_def)
apply (rule allI, rule impI)
 apply (simp add:Gcsrp_def)
 apply (rule gcsrp_l_one[of "N"], rule Group_axioms, assumption+)
done

lemma (in Group) gcsrp_map_gbijec:"G \<triangleright> N ==> 
                  gbij(G/N), (Gcsrp G N) (gcsrp_map G N)"
apply (simp add:gbijec_def)
apply (simp add:gsurjec_def ginjec_def)
apply (simp add:Qg_carrier Gcsrp_def) 
apply (frule nsg_sg[of "N"],
       frule gcsrp_map_bij[of "N"], simp add:bij_to_def)
apply (fold Gcsrp_def)
apply (simp add:gHom_def)

apply (rule conjI)
 apply (simp add:Qg_carrier gcsrp_map_def)
apply (rule conjI)
 apply (simp add:Qg_carrier Gcsrp_def) 
 apply (simp add:gcsrp_func1)
 apply (fold bij_to_def)
apply (rule ballI)+
 apply (simp add:Qg_def Gcsrp_def gcsrp_top_def)
 apply (frule gcsrp_func1[of "N"])
 apply (simp add:invfun_l1[of "gcsrp_map G N" "set_rcs G N" 
                                 "gcsrp_map G N ` set_rcs G N"])
done

lemma (in Group) Qg_equiv_Gcsrp:"G \<triangleright> N ==> (G / N) ≅ Gcsrp G N"
apply (simp add:isomorphic_def)
apply (frule gcsrp_map_gbijec[of "N"], blast)
done
 
section "14. Zassenhaus"

text{* we show H -> H N/N is gsurjective *}

lemma (in Group) homom4Tr1:"[|G \<triangleright> N; G » H|] ==>  Group ((Gp G (H \<struct>G N)) / N)" 
apply (frule Gp_smult_sg_nsg[of "H" "N"], assumption+)
apply (frule Gp_smult_nsg [of "H" "N"], assumption+)
 apply (simp add:Group.Group_Qg)
done

lemma homom3Tr2:"[|Group G; G » H; G \<triangleright> N|] ==>  
 gsurj(Gp G H),((Gp G (H \<struct>G N))/N) 
                      ((Pj (Gp G (H \<struct>G N)) N) o(Gp G H)(Gp G H)))"
 apply (frule iotahom[of "G" "H" "N"], assumption+,
        frule Group.Gp_smult_nsg[of "G" "H" "N"], assumption+,
        frule Group.smult_sg_nsg[of "G" "H" "N"], assumption+,
        frule Group.Gp_smult_sg_nsg[of "G" "H" "N"], assumption+,
        frule Group.Pj_gsurjec [of "Gp G (H \<struct>G N)" "N"], assumption,
        frule Group.Group_Gp[of "G" "H"], assumption+,
        frule Group.Group_Qg[of "Gp G (H \<struct>G N)" "N"], assumption+,
        frule gHomcomp[of "Gp G H" "Gp G (H \<struct>G N)" "(Gp G (H \<struct>G N)) / N" 
        "ι(\<natural>GH)" "Pj (Gp G (H \<struct>G N)) N"], assumption+) 
  apply (simp add:gsurjec_def)
 apply (subst gsurjec_def, simp)

 apply (rule surj_to_test,
        simp add:gHom_def)

 apply (rule ballI)
 apply (simp add:Group.Qg_carrier[of "Gp G (H \<struct>G N)" "N"],
        simp add:set_rcs_def, erule bexE,
        frule Group.nsg_sg[of "G" "N"], assumption,
        frule Group.r_sub_smult[of "G" "H" "N"], assumption+,
        simp add:Group.Gp_carrier)
 apply (simp add:Group.Gp_rcs[of "G" "N" "H \<struct>G N"])
 
 apply (thin_tac "ι(\<natural>GH) ∈ gHom (\<natural>GH) (\<natural>G(H \<struct>G N))",
        thin_tac "gsurj(Gp G (H \<struct>G N)),((\<natural>G(H \<struct>G N)) / N) Pj (\<natural>G(H \<struct>G N)) N",
        thin_tac "Pj (\<natural>G(H \<struct>G N)) N o(\<natural>GH) ι(\<natural>GH) ∈ gHom (\<natural>GH) ((\<natural>G(H \<struct>G N)) / N)")
 apply (simp add:cmpghom_def compose_def,
        simp add:Group.Gp_carrier iota_def,
        frule Group.smult_commute_sg_nsg[of "G" "H" "N"], assumption+,
        frule_tac a = a in eq_set_inc[of _ "H \<struct>G N" "N \<struct>G H"], assumption+,
        thin_tac "H \<struct>G N = N \<struct>G H")
 apply (simp add:s_top_def[of "G" "N" "H"], (erule bexE)+,
       rotate_tac -1, frule sym, thin_tac "x ·G y = a", 
       frule_tac h = y in Group.sg_subset_elem[of "G" "H"], assumption+,
       simp add:Group.rcs_fixed1[THEN sym])
 apply (frule Group.l_sub_smult[of "G" "H" "N"], assumption+,
        frule_tac x1 = y in Group.Pj_mem[THEN sym, of "Gp G (H \<struct>G N)" "N"],
        assumption+, simp add:Group.Gp_carrier, simp add: subsetD)
apply (frule_tac c = y in subsetD[of "H" "H \<struct>G N"], assumption+,
       simp add:Group.Gp_rcs[of "G" "N" "H \<struct>G N"], blast)
done


theorem homom4:"[|Group G; G \<triangleright> N; G » H|] ==>gbij((Gp G H)/(H ∩ N)),((Gp G (H \<struct>G N)) / N) (((Pj (Gp G (H \<struct>G N)) N) o(Gp G H)(Gp G H)))¨(Gp G H),((Gp G (H \<struct>G N)) / N))"
            
apply (frule homom3Tr2 [of "G" "H" "N"], assumption+)
apply (frule Group.Gp_smult_sg_nsg, assumption+)
apply (frule Group.homom4Tr1[of "G" "N" "H"], assumption+)
apply (frule Group.Group_Gp [of "G" "H"], assumption+)
apply (frule induced_ghom_ginjec [of "Gp G H" "(Gp G (H \<struct>G N)/N)" "(Pj (Gp G (H \<struct>G N)) N) o(Gp G H)(Gp G H))"], assumption+) 
 apply (simp add:gsurjec_def)
apply (frule inducedhomgsurjec [of "Gp G H" "(Gp G (H \<struct>G N))/N" "(Pj (Gp G (H \<struct>G N)) N) o(Gp G H)(Gp G H))"], assumption+)
 apply (frule Group.homom3Tr1[of "G" "H" "N"], assumption+)
 apply simp
apply (simp add:gbijec_def)   
done

lemma (in Group) homom4_2:"[|G \<triangleright> N; G » H|] ==> Group ((Gp G H) / (H ∩ N))" 
by (frule Group_Gp[of "H"],
    frule inter_Gp_nsg[of "N" "H"], assumption,
    rule Group.Group_Qg, assumption+)

lemma isom4:"[|Group G; G \<triangleright> N; G » H|] ==>
                 ((Gp G H)/(H ∩ N)) ≅  ((Gp G (N \<struct>G H)) / N)"
apply (frule homom4 [of "G" "N" "H"], assumption+,
       frule Group.smult_sg_nsg[of "G" "H" "N"], assumption+,
       frule Group.smult_commute_sg_nsg[of "G" "H" "N"], assumption+)
 apply (simp add:isomorphic_def, blast)
done

lemma ZassenhausTr5:"[|Group G; G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
      Gp G K \<triangleright> K1|] ==>
   ((Gp G (H ∩ K))/((H1 ∩ K) \<struct>G (H ∩ K1))) ≅ 
                          ((Gp G (H1 \<struct>G (H ∩ K)))/(H1 \<struct>G (H ∩ K1)))"
apply (frule Group.ZassenhausTr2_1 [of "G" "H" "H1" "K"], assumption+,
       frule Group.Group_Gp [of "G" "H1 \<struct>G (H ∩ K)"], assumption+,
       frule Group.ZassenhausTr3 [of "G" "H" "H1" "K" "K1"], assumption+,
       frule Group.inter_sgs [of "G" "H" "K"], assumption+,
       frule Group.r_sub_smult[of "G" "H1" "H ∩ K"], assumption+,
       frule Group.sg_sg[of "G" "H1 \<struct>G H ∩ K" "H ∩ K"], assumption+,
       frule isom4 [of "Gp G (H1 \<struct>G H ∩ K)" "H1 \<struct>G H ∩ K1" "H ∩ K"], 
                                                             assumption+)
apply (simp add:Int_commute[of "H ∩ K" "H1 \<struct>G H ∩ K1"])
apply (frule Group.Group_Gp[of "G" "H"], assumption,
       frule Group.Group_Gp[of "G" "K"], assumption,
       frule Group.nsg_sg[of "Gp G H" "H1"], assumption+,
       frule Group.sg_subset[of "Gp G H" "H1"], assumption+,
       frule Group.nsg_sg[of "Gp G K" "K1"], assumption+, 
       frule Group.sg_subset[of "Gp G K" "K1"], assumption+, 
       simp add:Group.Gp_carrier,
       frule Group.inter_sgs[of "G" "H" "K1"], assumption+,
       cut_tac subset_self[of "H"],
       frule Int_mono[of "H" "H" "K1" "K"], assumption)
apply (simp add:Group.s_topTr6_1[of "G" "H1" "H ∩ K1" "H ∩ K"],
       simp add:Int_assoc[THEN sym, of "H1" "H" "K"]) 
  
apply (simp add:Int_absorb2[of "H1" "H"],
       simp add:Group.Gp_inherited[of "G" "H ∩ K" "H1 \<struct>G H ∩ K"])
 apply (frule Group.s_top_mono[of "G" "H1" "H ∩ K" "H1" "H ∩ K1"],
        frule Group.sg_subset[of "G" "H"], assumption+,
        rule subset_trans[of "H1" "H" "carrier G"], assumption+)
  apply (rule Group.sg_subset[of "G" "H ∩ K"], assumption+, simp,
         simp,
         (frule Group.ZassenhausTr2_1[of "G" "H" "H1" "K"], assumption+,
          frule Group.subg_sg_sg[of "G" "H1 \<struct>G H ∩ K" "H1 \<struct>G H ∩ K1"],
                                  assumption+, simp add:Group.nsg_sg))
  apply (simp add:Group.s_top_induced[of "G" "H1 \<struct>G H ∩ K" "H1 \<struct>G H ∩ K1" "H ∩ K"],
         simp add:Group.s_top_assoc[of "G" "H1" "H ∩ K1" "H ∩ K"],
         cut_tac subset_self[of "H"],
         frule Int_mono[of "H" "H" "K1" "K"], assumption)
  apply (simp add:Group.K_absorb_HK[of "G" "H ∩ K1" "H ∩ K"])
  apply (cut_tac subset_self[of "H1 \<struct>G H ∩ K"],
         simp add:Group.Gp_inherited[of "G" "H1 \<struct>G H ∩ K" "H1 \<struct>G H ∩ K"])
done

lemma ZassenhausTr5_1:"[|Group G; G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1|] ==>   ((Gp G (K ∩ H))/((K1 ∩ H) \<struct>G (K ∩ H1))) ≅ 
                          ((Gp G (K1 \<struct>G (K ∩ H)))/(K1 \<struct>G (K ∩ H1)))"
(* thm ZassenhausTr5 [of "G" "K" "K1" "H" "H1"] *)
apply (simp add:ZassenhausTr5 [of "G" "K" "K1" "H" "H1"]) 
done

lemma ZassenhausTr5_2: "[|Group G; G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1|] ==>
      ((Gp G (H ∩ K))/((H1 ∩ K) \<struct>G (H ∩ K1))) = 
                       ((Gp G (K ∩ H))/((K1 ∩ H) \<struct>G (K ∩ H1)))"
by (simp add:Group.ZassenhausTr3_3[of "G" "H" "H1" "K" "K1"],
       simp add:Int_commute[of "H" "K"])

lemma ZassenhausTr6_1:"[|Group G; G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1|] ==> Group  (Gp G (H ∩ K) / (H1 ∩ K \<struct>G H ∩ K1))"
apply (frule Group.inter_sgs [of "G" "H" "K"], assumption+,
       frule Group.Group_Gp [of "G" "H ∩ K"], assumption+,
       frule Group.ZassenhausTr3_5 [of "G" "H" "H1" "K" "K1"], assumption+)
apply (rule Group.Group_Qg, assumption+)
done

lemma ZassenhausTr6_2:"[|Group G; G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1|] ==> Group (Gp G (H1 \<struct>G H ∩ K) / (H1 \<struct>G H ∩ K1))"
apply (frule Group.ZassenhausTr2_1 [of "G" "H" "H1" "K"], assumption+,
       frule Group.Group_Gp [of "G" "H1 \<struct>G H ∩ K"], assumption+,
       frule Group.ZassenhausTr3 [of "G" "H" "H1" "K" "K1"], assumption+)
apply (simp add:Group.Group_Qg)
done

lemma ZassenhausTr6_3:"[|Group G; G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1|] ==> Group (Gp G (K1 \<struct>G K ∩ H) / (K1 \<struct>G K ∩ H1))"
apply (frule Group.ZassenhausTr2_1 [of "G" "K" "K1" "H"], assumption+,
       frule Group.Group_Gp [of "G" "K1 \<struct>G K ∩ H"], assumption+,
       frule Group.ZassenhausTr3[of "G" "K" "K1" "H" "H1"], assumption+)
apply (simp add:Group.Group_Qg)
done

theorem Zassenhaus:"[|Group G; G » H; G » H1; G » K; G » K1; Gp G H \<triangleright> H1; 
       Gp G K \<triangleright> K1|] ==> (Gp G (H1 \<struct>G H ∩ K) / (H1 \<struct>G H ∩ K1)) ≅ 
                              (Gp G (K1 \<struct>G K ∩ H) / (K1 \<struct>G K ∩ H1))" 
apply (frule ZassenhausTr6_1[of "G" "K" "K1" "H" "H1"], assumption+)
apply (frule ZassenhausTr6_3 [of "G" "H" "H1" "K" "K1"], assumption+)
apply (frule ZassenhausTr6_2 [of "G" "H" "H1" "K" "K1"], assumption+)
apply (rule isomTr2[of "(\<natural>G(H1 \<struct>G H ∩ K)) / (H1 \<struct>G H ∩ K1)" 
             "(\<natural>G(K ∩ H)) / (K1 ∩ H \<struct>G K ∩ H1)"
             "(\<natural>G(K1 \<struct>G K ∩ H)) / (K1 \<struct>G K ∩ H1)"], assumption+)
apply (frule ZassenhausTr5_1[of "G" "K" "K1" "H" "H1"], assumption+)
apply (simp add:Int_commute[of "K" "H"])
apply (simp add:Group.ZassenhausTr3_3[THEN sym, of "G" "H" "H1" "K" "K1"])
apply (rule isomTr1[of  "(\<natural>G(H ∩ K)) / (H1 ∩ K \<struct>G H ∩ K1)" 
                        "(\<natural>G(H1 \<struct>G H ∩ K)) / (H1 \<struct>G H ∩ K1)"], assumption+)
apply (rule ZassenhausTr5_1[of "G" "H" "H1" "K" "K1"], assumption+)
done


section "15. chain of groups I"  

constdefs (structure G)
  d_gchain:: "[_ , nat, (nat => 'a set)] => bool"
    "d_gchain G n g  == if n=0 then G » g 0 else (∀l≤ n. G » (g l) ∧ 
            (∀l ≤ (n - Suc 0). g (Suc l) ⊆ g l ))" 
            (**  g 0 ⊇ … ⊇ g n  **)

constdefs (structure G)
  D_gchain:: "[_ , nat, (nat => 'a set)] => bool"
  "D_gchain G n g == if n = 0 then G » (g 0) else (d_gchain G n g) ∧ 
      (∀l ≤ (n - Suc 0). (g (Suc l)) ⊂ (g l))"
            (**  g 0 ⊃ … ⊃ g n **) 

constdefs (structure G)
 td_gchain::"[_ , nat, (nat => 'a set)] => bool"
 "td_gchain G n g == if n=0 then g 0 = carrier G ∧ g 0 = {\<one>} else 
       d_gchain G n g ∧ g 0 = carrier G ∧ g n = {\<one>}"
           (**  g 0 ⊇ … ⊇ g n with g 0 = carrier G and g n = {e}  **)
constdefs (structure G)
 tD_gchain::"[_, nat, (nat => 'a set)] => bool"
 "tD_gchain G n g == if n=0 then g 0 = carrier G ∧ g 0 = {\<one>} else 
        D_gchain G n g ∧ (g 0 = carrier G) ∧ (g n = {\<one>})" 
           (**  g 0 ⊃ … ⊃ g n with g 0 = carrier G and g n = {e}  **)

 w_cmpser::"[_ , nat, (nat => 'a set)] => bool"
  "w_cmpser G n g == if n = 0 then d_gchain G n g else d_gchain G n g ∧  
    (∀l ≤ (n - 1). (Gp G (g l)) \<triangleright> (g (Suc l)))" 
           (**  g 0 \<rhd> … \<rhd> g n ** with g l ⊇ g (Suc l) **) 


 W_cmpser::"[_ , nat, (nat => 'a set)] => bool"
  "W_cmpser G n g == if n = 0 then d_gchain G 0 g else D_gchain G n g ∧ 
  (∀l ≤ (n - 1). (Gp G (g l)) \<triangleright> (g (Suc l)))" 
           (**  g 0 \<rhd> … \<rhd> g n  with g i ⊃ g (Suc i) **) 

 tw_cmpser::"[_ , nat, (nat => 'a set)] => bool"
  "tw_cmpser G n g == if n = 0 then td_gchain G 0 g else td_gchain G n g ∧ 
  (∀l ≤ (n - 1). (Gp G (g l)) \<triangleright> (g (Suc l)))" 
         (**  g 0 \<rhd> … \<rhd> g n with g 0 = carrier G and g n = {e}  **) 

 tW_cmpser::"[_ , nat, (nat => 'a set)] => bool"
 "tW_cmpser G n g == if n = 0 then td_gchain G 0 g else tD_gchain G n g ∧ 
     (∀l ≤ (n - 1). (Gp G (g l)) \<triangleright> (g (Suc l)))"
    (* g 0 \<rhd> … \<rhd> g n with g 0 = carrier G, g n = {e} and g (Suc i) ⊂ g i*) 

 Qw_cmpser::"[_ , nat => 'a set] => 
                                (nat => ('a set) Group)" 
  "Qw_cmpser G f l == ((Gp G (f l)) / (f (Suc l)))" 

 (* f 0 \<rhd> … \<rhd> f (n+1), Qw_cmpser G n f is (f 0)/(f 1),…,f n/f (n+1) *)

constdefs (structure G)
 red_chn:: "[_ , nat, (nat => 'a set)] => (nat => 'a set)"  
 "red_chn G n f == SOME g. g ∈ {h.(tW_cmpser G (card (f ` {i. i ≤ n}) - 1) h)
     ∧ h `{i. i ≤ (card (f ` {i. i ≤ n}) - 1)} = f `{i. i ≤ n}}"

  chain_cutout :: "[nat, (nat => 'a set) ]
                                 => (nat => 'a set)"
      "chain_cutout l f  == λj. f (slide l j)"

lemma (in Group) d_gchainTr0:"[|0 < n; d_gchain G n f; k ≤ (n - 1)|]
                        ==> f (Suc k) ⊆ f k"
apply (simp add:d_gchain_def) 
apply (frule_tac a = k in forall_spec)
 apply (rule Nat.le_trans, assumption+, simp)
 apply (erule conjE, rotate_tac 2,
       frule_tac a = k in forall_spec, assumption, 
       thin_tac "∀l≤n - Suc 0. f (Suc l) ⊆ f l",
       thin_tac "∀l≤n. G » f l  ∧ (∀l≤n - Suc 0. f (Suc l) ⊆ f l)")
 apply assumption
done

lemma (in Group) d_gchain_mem_sg:"d_gchain G n f ==> ∀i≤ n. G » (f i)"
apply (rule allI)
apply (rule impI, simp add:d_gchain_def)
apply (case_tac "n = 0", simp)
apply simp
done

lemma (in Group) d_gchain_pre:"d_gchain G (Suc n) f ==> d_gchain G n f"
apply (simp add:d_gchain_def, rule impI, rule impI)
apply (rule allI, rule impI)
apply (frule_tac a = l in forall_spec, arith)
apply (erule conjE)
apply (thin_tac "∀l≤Suc n. G » f l  ∧ (∀l≤n. f (Suc l) ⊆ f l)",
       frule_tac a = l in forall_spec, arith, assumption) 
done

lemma (in Group) d_gchainTr1:"0 < n --> (∀f. d_gchain G n f --> 
                     (∀l ≤ n. ∀j ≤ n. l < j --> f j ⊆ f l))"
apply (induct_tac n)
 apply (rule impI, simp) 
(** n **)
apply (rule impI, rule allI, rule impI)
apply ((rule allI, rule impI)+, rule impI)
(** case n = 0 **)
apply (case_tac "n = 0", simp)
 apply (case_tac "j = 0", simp, 
        frule le_imp_less_or_eq, thin_tac "j ≤ Suc 0",
        simp, simp add:d_gchain_def)
 apply (frule_tac a = 0 in forall_spec, simp, simp)

(** case 0 < n **) 
 apply simp
 (** case j = n **)
apply (case_tac "j = Suc n")
 apply (frule d_gchain_pre, 
        frule_tac a = f in forall_spec, assumption,
        thin_tac "∀f. d_gchain G n f --> (∀l≤n. ∀j≤n. l < j --> f j ⊆ f l)",
        thin_tac "d_gchain G n f",
        simp add:d_gchain_def)
 apply (frule_tac a = n in forall_spec, simp,
        thin_tac "∀l≤Suc n. G » f l  ∧ (∀l≤n. f (Suc l) ⊆ f l)",
        erule conjE,
        rotate_tac -1,
        frule_tac a = n in forall_spec, simp,
        thin_tac "∀l≤n. f (Suc l) ⊆ f l",
        frule_tac x = l and n = n in Suc_less_le)
  apply (case_tac "l = n", simp,
         thin_tac "l < Suc n",
         frule_tac x = l and y = n in le_imp_less_or_eq, 
         thin_tac "l ≤ n", simp)
  apply (frule_tac a = l in forall_spec, simp,
         thin_tac "∀l≤n. ∀j≤n. l < j --> f j ⊆ f l") apply (
         frule_tac a = n in forall_spec) apply (simp,
         thin_tac "∀j≤n. l < j --> f j ⊆ f l", simp)
  apply (metis less_le_trans d_gchain_pre less_Suc_eq_le linorder_antisym_conv1 mod_if mod_le_divisor)
done

lemma (in Group) d_gchainTr2:"[|0 < n; d_gchain G n f; l ≤ n; j ≤ n; l ≤ j |]
                               ==> f j ⊆ f l"
apply (case_tac "l = j", simp)
apply (metis Group.d_gchainTr1 Group_axioms linorder_antisym_conv2)
done

lemma (in Group) im_d_gchainTr1:"[|d_gchain G n f;
       f l ∈ (f ` {i. i ≤ n}) - {f 0}|] ==> 
    f (LEAST j. f j ∈ (f ` {i. i ≤ n}) - {f 0}) ∈ (f ` {i. i ≤ n} - {f 0})"
apply (rule LeastI)
apply simp
done

lemma (in Group) im_d_gchainTr1_0:"[|d_gchain G n f; 
                 f l ∈ (f ` {i. i ≤ n}) - {f 0}|]  ==> 
                 0 < (LEAST j. f j ∈ (f ` {i. i ≤ n}) - {f 0})"
apply (frule im_d_gchainTr1 [of "n" "f"], assumption+)
apply (rule contrapos_pp, simp)
apply simp
done 

lemma (in Group) im_d_gchainTr1_1:
      "[|d_gchain G n f; ∃ i. f i ∈ (f ` {i. i ≤ n}) - {f 0}|]  ==>
  f (LEAST j. f j ∈ ((f ` {i. i ≤ n}) - {f 0})) ∈ ((f` {i. i ≤ n}) - {f 0})"
apply (subgoal_tac "∀i. f i ∈ f ` {i. i ≤ n} - {f 0} --> 
        f (LEAST j. f j ∈ f `{i. i ≤ n} - {f 0}) ∈ f ` {i. i≤ n} - {f 0}")
apply blast
 apply (thin_tac "∃i. f i ∈ f `{i. i ≤ n} - {f 0}")
 apply (rule allI) apply (rule impI)
apply (rule im_d_gchainTr1 [of "n" "f" _], assumption+)
done

lemma (in Group) im_d_gchainsTr1_2:"
      [|d_gchain G n f; i ≤ n; f i ∈ f `{i. i ≤ n} - {f 0}|]  ==>
          (LEAST j. f j ∈ (f `{i. i ≤ n} - {f 0})) ≤ i"
by (rule Least_le, assumption)

lemma (in Group) im_d_gchainsTr1_3:"[|d_gchain G n f; ∃i ≤ n. 
                 f i ∈ f` {i. i ≤ n} - {f 0};
                 k < (LEAST j. f j ∈ (f `{i. i ≤ n} - {f 0}))|] ==> f k = f 0"
apply (erule exE)
apply (rule contrapos_pp, simp+)
apply (frule_tac i = i in im_d_gchainsTr1_2 [of "n" "f" _ ], simp+)  
apply (erule conjE)+ 
apply (frule_tac x = k and 
                 y = "LEAST j. f j ∈ f ` {i. i ≤ n} ∧ f j ≠ f 0" and
                 z = i in less_le_trans, assumption,
       frule_tac x = k and 
                 y = i and
                 z = n in less_le_trans, assumption)
apply (frule_tac i = k in im_d_gchainsTr1_2 [of "n" "f" _ ], simp+)  
done
 
lemma (in Group) im_gdchainsTr1_4:"[|d_gchain G n f;
       ∃v∈f `{i. i ≤ n}. v ∉ {f 0}; i < (LEAST j. f j ∈ (f `{i. i ≤ n}) ∧ 
       f j ≠ f 0) |] ==> f i = f 0"
apply (rule im_d_gchainsTr1_3 [of "n" "f" "i"], assumption,
       thin_tac "i < (LEAST j. f j ∈ f ` {i. i ≤ n} ∧ f j ≠ f 0)",
       simp add:image_def, blast)
apply simp
done

lemma (in Group) im_d_gchainsTr1_5:"[|0 < n; d_gchain G n f; i ≤ n; 
f i ∈ (f ` {i. i ≤ n} - {f 0}); (LEAST j. f j ∈ (f `{i. i ≤ n} - {f 0})) = j|] 
==>         f `{i. i ≤ (j - (Suc 0))} = {f 0}"
apply (frule im_d_gchainTr1_0 [of "n" "f" "i"], assumption+)
apply (subst image_def)
apply (rule equalityI)
apply (rule subsetI, simp, erule exE, erule conjE)
 apply (frule_tac x = xa and y = "j - Suc 0" and 
        z = "(LEAST j. f j ∈ f ` {i. i ≤ n} ∧ f j ≠ f 0)" in le_less_trans,
        simp,
        frule_tac k = xa in im_d_gchainsTr1_3[of "n" "f"])
  apply (simp, blast, simp, simp) 
apply (rule subsetI, blast)
done

lemma (in Group) im_d_gchains1:"[|0 < n; d_gchain G n f; i ≤ n; 
                 f i ∈ (f ` {i. i ≤ n} - {f 0}); 
                 (LEAST j. f j ∈ (f `{i. i ≤ n} - {f 0})) = j |] ==> 
                         f `{i. i ≤ n} = {f 0} ∪ {f i |i. j ≤ i ∧ i ≤ n}"
apply (frule im_d_gchainTr1_0 [of "n" "f" "i"], assumption+,
       frule im_d_gchainsTr1_2 [of "n" "f" "i"], assumption+,
       frule Nset_nset_1 [of "n" "j - Suc 0"])
apply simp;
apply (subst im_set_un2, simp)
apply (subst im_d_gchainsTr1_5[of "n" "f" "i" "j"]) 
apply (simp, assumption, simp+)

apply (rule equalityI)
 apply (rule subsetI, simp, erule disjE, simp,
        simp add:image_def nset_def, erule exE, blast) 
 apply (rule subsetI)
 apply (simp, erule disjE, simp)
 apply (erule exE, simp add:nset_def)
done

lemma (in Group) im_d_gchains1_1:"[|d_gchain G n f; f n ≠ f 0|]  ==> 
      f `{i. i ≤ n} = {f 0} ∪ 
         {f i |i. (LEAST j. f j ∈ (f `{i. i ≤ n} - {f 0})) ≤ i ∧ i ≤ n}"
apply (case_tac "n = 0") apply simp
apply (simp)
apply (frule im_d_gchains1 [of "n" "f" "n" 
                "(LEAST j. f j ∈ f ` {i. i ≤ n} - {f 0})"], assumption+,
       simp add:n_in_Nsetn)
apply (cut_tac n_in_Nsetn[of "n"], simp,
       simp)
apply simp
done

lemma (in Group) d_gchains_leastTr:"[|d_gchain G n f; f n ≠ f 0|]  ==>  
               (LEAST j. f j ∈ (f `{i. i ≤ n} - {f 0})) ∈ {i. i ≤ n} ∧ 
               f  (LEAST j. f j ∈ (f `{i. i ≤ n} - {f 0})) ≠ f 0"
apply (frule im_d_gchainsTr1_2 [of "n" "f" "n"],
       simp add:n_in_Nsetn,
       simp add:image_def, blast,
       frule mem_of_Nset[of "LEAST j. f j ∈ f ` {i. i ≤ n} - {f 0}" "n"],
       simp)
apply (frule im_d_gchainTr1[of "n" "f" "n"],
       simp add:image_def, cut_tac n_in_Nsetn[of "n"], blast)
apply (simp add:image_def)
done

lemma (in Group) im_d_gchainTr2:"[|d_gchain G n f; j ≤ n; f j ≠ f 0|] ==>
                                 ∀i ≤ n. f 0 = f i --> ¬ j ≤ i"
apply (case_tac "n = 0", simp, simp)
apply (rule allI, rule impI)
apply (rule contrapos_pp, simp+)
apply (case_tac "j = i", simp) 
apply (frule d_gchainTr2 [of "n" "f" "0" "j"], assumption+,
       simp, (erule conjE)+,
       rule_tac i = j and j = i and k = n in le_trans, assumption+,
       simp,
       (erule conjE)+,
       frule_tac l = j and j = i in d_gchainTr2 [of "n" "f"], assumption+)

       apply simp+
done

lemma (in Group) D_gchain_pre:"[|D_gchain G (Suc n) f|] ==> D_gchain G n f"
apply (simp add:D_gchain_def, erule conjE)
apply (case_tac "n = 0", rotate_tac -1)
 apply (simp, insert lessI [of "0::nat"], simp)
 apply (simp add:d_gchain_def, simp)
 apply (frule d_gchain_pre [of "n"])
 apply simp 
done

lemma (in Group) D_gchain0:"[|D_gchain G n f; i ≤ n; j ≤ n; i < j|] ==>
                             f j ⊂ f i" 
apply (case_tac "n = 0") 
 apply (simp, simp)
apply (cut_tac d_gchainTr1[of "n"], simp)
 apply (simp add:D_gchain_def, frule conjunct1)
 apply (frule_tac a = f in forall_spec, assumption,
        thin_tac "∀f. d_gchain G n f --> (∀l≤n. ∀j≤n. l < j --> f j ⊆ f l)") 
 apply (frule_tac a = i in forall_spec,
        frule_tac x = i and y = j and z = n in less_le_trans, assumption+,
        simp)
apply ( thin_tac "∀l≤n. ∀j≤n. l < j --> f j ⊆ f l",
        frule_tac a = j in forall_spec, assumption, 
        thin_tac "∀j≤n. i < j --> f j ⊆ f i", simp) 
apply (frule Suc_leI[of i j],
       frule less_le_trans[of i j n], assumption+,
       frule less_le_diff[of i n])
apply (frule_tac a = i in forall_spec, assumption,
       thin_tac "∀l≤n - Suc 0. f (Suc l) ⊂ f l")
 apply (cut_tac d_gchainTr2[of "n" "f" "Suc i" "j"])
 apply blast apply simp+
done

lemma (in Group) D_gchain1:"D_gchain G n f ==> inj_on f {i. i ≤ n}"
apply (case_tac "n = 0")
 apply (simp add:inj_on_def)
apply (simp) (** case 0 < n **)
apply (simp add:inj_on_def)
 apply ((rule allI)+, rule impI, rule contrapos_pp, simp+, erule exE)
 apply (cut_tac x = x and y = y in less_linear, simp)
apply (erule disjE, (erule conjE)+)
 apply (frule_tac i = x and j = y in D_gchain0 [of "n" "f"], assumption+,
        simp,
        simp, frule_tac i = y and j = x in D_gchain0 [of "n" "f"],
        simp+)
done

lemma (in Group) card_im_D_gchain:"[|0 < n; D_gchain G n f|] 
                                   ==> card (f `{i. i ≤ n}) = Suc n"
apply (insert finite_Nset [of "n"], frule D_gchain1 [of "n"])
apply (subst card_image, assumption+, simp add:card_Nset)
done

lemma (in Group) w_cmpser_gr:"[|0 < r; w_cmpser G r f; i ≤ r|]
                             ==> G »  (f i)"
by (simp add:w_cmpser_def, erule conjE, simp add:d_gchain_def)

lemma (in Group) w_cmpser_ns:"[|0 < r; w_cmpser G r f; i ≤ (r - 1)|] ==>
                                 (Gp G (f i)) \<triangleright> (f (Suc i))"
apply (simp add:w_cmpser_def)
done

lemma (in Group) w_cmpser_pre:"w_cmpser G (Suc n) f ==> w_cmpser G n f"
apply (simp add:w_cmpser_def)
 apply (erule conjE)
 apply (case_tac "n = 0", rotate_tac -1, simp)
 apply (rule d_gchain_pre [of "0" "f"], assumption+)
apply simp
 apply (frule d_gchain_pre [of "n" "f"])
 apply simp
done

lemma (in Group) W_cmpser_pre:"W_cmpser G (Suc n) f ==> W_cmpser G n f"
apply (simp add:W_cmpser_def)
 apply (erule conjE)
 apply (case_tac "n = 0", simp,
        simp add:D_gchain_def, erule conjE,
        rule d_gchain_pre, assumption+, simp)

apply (frule D_gchain_pre, simp);
done

lemma (in Group) td_gchain_n:"[|td_gchain G n f; carrier G ≠ {\<one>}|] ==> 0 < n"
apply (simp add:td_gchain_def)
apply (rule contrapos_pp, simp+) 
apply (erule conjE, simp) 
done

section "16. Existence of reduced chain" 

lemma (in Group) D_gchain_is_d_gchain:"D_gchain G n f ==> d_gchain G n f"
apply (simp add:D_gchain_def)
 apply (case_tac "n = 0") apply (rotate_tac -1) 
 apply (simp add:d_gchain_def) apply (rotate_tac -1)
 apply simp
done

lemma (in Group) joint_d_gchains:"[|d_gchain G n f; d_gchain G m g; 
  g 0 ⊆ f n |] ==>  d_gchain G (Suc (n + m)) (jointfun n f m g)" 
apply (case_tac "n = 0")
 apply (case_tac "m = 0")
 apply (simp add:d_gchain_def [of "G" "Suc (n + m)" _])
 apply (simp add:jointfun_def sliden_def d_gchain_def)
 apply (simp add:jointfun_def sliden_def d_gchain_def)
 apply (rule allI) apply (rule conjI) apply (rule impI) 
 apply (rule allI) apply (rule impI)+ apply simp 
 apply (frule_tac a = la in forall_spec, assumption,
        thin_tac "∀l≤m. G » g l  ∧ (∀l≤m - Suc 0. g (Suc l) ⊆ g l)",
        erule conjE)
 apply (frule_tac a = "la - Suc 0" in forall_spec,
        thin_tac "∀l≤m - Suc 0. g (Suc l) ⊆ g l",
        rule diff_le_mono, assumption, simp)
 apply (rule impI, rule impI) 
 apply (frule_tac m = l and n = "Suc m" and l = "Suc 0" in diff_le_mono)
       apply simp
       apply (rule allI, rule impI, rule impI) 
       apply (frule_tac a = la in forall_spec,assumption,
              thin_tac "∀l≤m. G » g l  ∧ (∀l≤m - Suc 0. g (Suc l) ⊆ g l)",
              erule conjE)
        apply (frule_tac a = "la - Suc 0" in forall_spec, simp) 
apply simp_all
 apply (simp add:d_gchain_def [of "G" _ "jointfun n f m g"])
 apply (rule allI, rule impI) apply (rule conjI)
 apply (case_tac "l ≤ n", simp add:jointfun_def d_gchain_def[of _ n f])
 apply (simp add:jointfun_def sliden_def,
        frule_tac m = l and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono,
        thin_tac "l ≤ Suc (n + m)", simp add:d_gchain_def[of _ m g])
 apply (case_tac "m = 0", simp, simp)

apply (rule allI, rule impI)
 apply (case_tac "Suc la ≤ n")
 apply (simp add:jointfun_def)
 apply (rule_tac l = la and j = "Suc la" in d_gchainTr2[of n f],
        simp+)
 apply (simp add:jointfun_def)
 apply (cut_tac y = "Suc la" and x = n in not_less [symmetric], simp)
 apply (frule_tac m = n and n = "Suc la" in Suc_leI,
        thin_tac "n < Suc la", simp)
 apply (case_tac "la = n", simp add:sliden_def)
apply (frule not_sym, thin_tac "la ≠ n", 
       frule_tac x = n and y = la in le_imp_less_or_eq, 
       thin_tac "n ≤ la", simp,
       frule_tac m = n and n = la in Suc_leI, simp add:sliden_def)
 apply (simp add:d_gchain_def[of _ m g])
 apply (cut_tac m = la and n = "n + m" and l = "Suc n" in diff_le_mono,
         assumption, simp)
 apply (frule_tac a = m in forall_spec, simp,
        thin_tac "∀l≤m. G » g l  ∧ (∀l≤m - Suc 0. g (Suc l) ⊆ g l)",
        erule conjE) apply (
        frule_tac a = "la - Suc n" in forall_spec, assumption,
        thin_tac "∀l≤m - Suc 0. g (Suc l) ⊆ g l")
 apply (cut_tac n1 = n and i1 = la  in jointgd_tool4[THEN sym], simp+)
done
 
lemma (in Group) joint_D_gchains:"[|D_gchain G n f; D_gchain G m g; 
       g 0 ⊂ f n |] ==>  D_gchain G (Suc (n + m)) (jointfun n f m g)" 
apply (simp add:D_gchain_def [of "G" "Suc (n + m)" _])
apply (rule conjI)
apply (rule joint_d_gchains[of n f m g],
       simp add:D_gchain_is_d_gchain,
       simp add:D_gchain_is_d_gchain,
       simp add:psubset_imp_subset)
apply (rule allI, rule impI)
 apply (case_tac "Suc l ≤ n")
 apply (simp add:jointfun_def)
 apply (rule_tac i = l and j = "Suc l" in D_gchain0[of n f], assumption,
        cut_tac x = l and y = "Suc l" and z = n in less_le_trans) 
        apply simp+
 apply (simp add:nat_not_le_less,
        frule_tac m = n and n = "Suc l" in Suc_leI, thin_tac "n < Suc l", simp)
 apply (case_tac "l = n", simp add:jointfun_def sliden_def)
 apply (frule not_sym, thin_tac "l ≠ n",
        frule_tac x = n and y = l in le_imp_less_or_eq,
        thin_tac "n ≤ l", simp)
 apply (simp add:jointfun_def sliden_def)
 apply (frule_tac m = l and n = "n + m" and l = n in  diff_le_mono)
        apply (simp add:diff_add_assoc)
 apply (rule_tac i = "l - Suc n" and j = "l - n" in D_gchain0[of m g],
         assumption) 
 apply (arith, assumption, arith)
done
  
lemma (in Group) w_cmpser_is_d_gchain:"w_cmpser G n f ==> d_gchain G n f"
apply (simp add:w_cmpser_def)
 apply (case_tac "n=0") apply (rotate_tac -1) apply simp
 apply (rotate_tac -1) apply simp
done

lemma (in Group) joint_w_cmpser:"[|w_cmpser G n f; w_cmpser G m g; 
 Gp G (f n) \<triangleright> (g 0)|] ==> w_cmpser G (Suc (n + m)) (jointfun n f m g)"
apply (simp add:w_cmpser_def [of _ "Suc (n + m)" _])
apply (rule conjI)
 apply (frule w_cmpser_is_d_gchain[of "n" "f"],
        frule w_cmpser_is_d_gchain[of "m" "g"])
  apply (rule joint_d_gchains, assumption+) 
  apply (frule d_gchain_mem_sg[of "n" "f"],
         cut_tac n_in_Nsetn[of "n"],
         frule_tac a = n in forall_spec, simp,
         thin_tac "∀i ≤ n. G » f i")
  apply (frule Group_Gp[of "f n"],
         frule Group.nsg_sg[of "Gp G (f n)" "g 0"], assumption,
         frule Group.sg_subset[of "Gp G (f n)" "g 0"], assumption,
         simp add:Gp_carrier)
apply (rule allI, rule impI)
 apply (case_tac "n = 0") apply simp
 apply (simp add:jointfun_def)
  apply (case_tac "l = 0")
  apply simp apply (simp add:sliden_def)
 apply simp
 apply (simp add:w_cmpser_def [of _ "m" "g"])
 apply (case_tac "m = 0") apply (simp add:sliden_def)
 apply (erule conjE)
 apply (simp add:sliden_def)
 apply (frule_tac x = 0 and y = l and z = m in less_le_trans, assumption+) 
 apply (frule_tac m = l and n = m and l = "Suc 0" in diff_le_mono)
 apply (frule_tac a = "l - Suc 0" in forall_spec, assumption,
        thin_tac "∀l≤(m - Suc 0). (\<natural>(g l)) \<triangleright> (g (Suc l))")
 apply simp
 apply (case_tac "l ≤ n - Suc 0", simp) 
 apply (frule less_pre_n [of "n"])
 apply (frule_tac x = l in le_less_trans[of _ "n - Suc 0" "n"], assumption+) 
 apply (simp add:jointfun_def w_cmpser_def [of _ "n"])
apply (simp add:nat_not_le_less)
 apply (frule_tac n = l in Suc_leI [of "n - Suc 0" _], simp)
 apply (case_tac "n = l")
 apply (frule sym) apply (thin_tac "n = l")
 apply simp
 apply (simp add:jointfun_def sliden_def)
 apply (frule_tac m = n and n = l in noteq_le_less, assumption+)
 apply (frule_tac m = n and n = l in Suc_leI)
 apply (simp add:jointfun_def)
 apply (frule_tac m = l and n = "n + m" and l = "Suc n" in diff_le_mono)
 apply simp
 apply (simp add:sliden_def w_cmpser_def [of _ "m" _])
 apply (erule conjE)
 apply (simp add:jointgd_tool4)
done  

lemma (in Group) W_cmpser_is_D_gchain:"W_cmpser G n f ==> D_gchain G n f"
apply (simp add:W_cmpser_def)
 apply (case_tac "n = 0") apply (rotate_tac -1) apply simp
 apply (simp add:D_gchain_def d_gchain_def)
 apply (rotate_tac -1) apply simp
done

lemma (in Group) W_cmpser_is_w_cmpser:"W_cmpser G n f ==> w_cmpser G n f"
apply (simp add:W_cmpser_def)
apply (case_tac "n = 0") apply (rotate_tac -1)
 apply simp
 apply (simp add:w_cmpser_def)
 apply (rotate_tac -1)
apply simp apply (erule conjE)
 apply (frule D_gchain_is_d_gchain)
 apply (simp add:w_cmpser_def)
done

lemma (in Group) tw_cmpser_is_w_cmpser:"tw_cmpser G n f ==> w_cmpser G n f"
apply (simp add:tw_cmpser_def)
 apply (case_tac "n = 0")
 apply (rotate_tac -1) apply simp
 apply (simp add:td_gchain_def w_cmpser_def) 
 apply (simp add:d_gchain_def) apply (simp add:special_sg_G)
apply (rotate_tac -1) apply simp
 apply (erule conjE) apply (simp add:td_gchain_def)
 apply (erule conjE)+
apply (simp add:w_cmpser_def)
done

lemma (in Group) tW_cmpser_is_W_cmpser:"tW_cmpser G n f ==> W_cmpser G n f"
apply (simp add:tW_cmpser_def)
apply (case_tac "n = 0") apply (rotate_tac -1)
 apply simp
 apply (simp add:td_gchain_def)
 apply (simp add:W_cmpser_def d_gchain_def) apply (simp add:special_sg_G)
apply (rotate_tac -1) apply simp
apply (erule conjE)
 apply (simp add:tD_gchain_def)
 apply (erule conjE)+
 apply (simp add:W_cmpser_def)
done

lemma (in Group) joint_W_cmpser:"[|W_cmpser G n f; W_cmpser G m g; 
      (Gp G (f n)) \<triangleright> (g 0); g 0 ⊂ f n|] ==> 
      W_cmpser G (Suc (n + m)) (jointfun n f m g)"
apply (simp add:W_cmpser_def [of _ "Suc (n + m)" _])
 apply (frule W_cmpser_is_D_gchain [of "n" "f"],
        frule W_cmpser_is_D_gchain [of "m" "g"])
 apply (simp add:joint_D_gchains)
apply (frule W_cmpser_is_w_cmpser [of "n" _],
       frule W_cmpser_is_w_cmpser [of "m" _])
 apply (frule joint_w_cmpser [of "n" "f" "m" "g"], assumption+)
 apply (simp add:w_cmpser_def [of _ "Suc (n + m)" _])
done

lemma (in Group) joint_d_gchain_n0:"[|d_gchain G n f; d_gchain G 0 g;
       g 0 ⊆ f n |] ==>  d_gchain G (Suc n) (jointfun n f 0 g)"
apply (frule joint_d_gchains [of "n" "f" "0" "g"], assumption+)
apply simp
done

lemma (in Group) joint_D_gchain_n0:"[|D_gchain G n f; D_gchain G 0 g; 
       g 0 ⊂ f n |] ==>  D_gchain G (Suc n) (jointfun n f 0 g)" 
apply (frule joint_D_gchains [of "n" "f" "0" "g"], assumption+)
apply simp
done

lemma (in Group) joint_w_cmpser_n0:"[|w_cmpser G n f; w_cmpser G 0 g; 
      (Gp G (f n)) \<triangleright> (g 0)|] ==>  w_cmpser G (Suc n) (jointfun n f 0 g)" 
apply (frule joint_w_cmpser [of "n" "f" "0" "g"], assumption+)
apply simp
done

lemma (in Group) joint_W_cmpser_n0:"[|W_cmpser G n f; W_cmpser G 0 g; 
      (Gp G (f n)) \<triangleright> (g 0); g 0 ⊂ f n |] ==>  
                       W_cmpser G (Suc n) (jointfun n f 0 g)" 
apply (frule joint_W_cmpser [of "n" "f" "0" "g"], assumption+)
apply simp
done

constdefs (structure G)
 simple_Group :: "_ => bool"
    "simple_Group G == {N. G » N} = {carrier G, {\<one>}}"

 compseries:: "[_ , nat, nat => 'a set] => bool"
   "compseries G n f == tW_cmpser G n f ∧ (if n = 0 then f 0 = {\<one>} else 
   (∀ i ≤ (n - 1). (simple_Group ((Gp G (f i))/(f (Suc i))))))"

 length_twcmpser:: "[_ , nat, nat => 'a set] => nat"
   "length_twcmpser G n f == card (f `{i. i ≤ n}) - Suc 0" 


lemma (in Group) compseriesTr0:"[|compseries G n f; i ≤ n|] ==>
                                    G » (f i)"
apply (simp add:compseries_def) 
 apply (frule conjunct1)
 apply (fold compseries_def)
 apply (frule tW_cmpser_is_W_cmpser,
        frule W_cmpser_is_w_cmpser, 
        frule w_cmpser_is_d_gchain)
apply (frule d_gchain_mem_sg[of "n" "f"])
 apply simp
done

lemma (in Group) compseriesTr1:"compseries G n f ==> tW_cmpser G n f"
apply (simp add:compseries_def)
done

lemma (in Group) compseriesTr2:"compseries G n f ==> f 0 = carrier G"
apply (frule compseriesTr1, simp add:tW_cmpser_def)
apply (case_tac "n = 0")
 apply (simp add:td_gchain_def)
apply simp
apply (erule conjE, simp add:tD_gchain_def)
done

lemma (in Group) compseriesTr3:"compseries G n f ==> f n = {\<one>}"
apply (frule compseriesTr1)
apply (simp add:tW_cmpser_def)
apply (case_tac "n = 0")
apply (simp add:td_gchain_def)
apply (auto del:equalityI)
apply (simp add:tD_gchain_def)
done

lemma (in Group) compseriesTr4:"compseries G n f ==> w_cmpser G n f"
apply (frule compseriesTr1,
       frule tW_cmpser_is_W_cmpser)
apply (rule W_cmpser_is_w_cmpser, assumption)
done

lemma (in Group) im_jointfun1Tr1:"∀l ≤ n. G » (f l) ==>
                               f ∈ {i. i ≤ n} -> Collect (sg G)"
apply (simp add:Pi_def)
done

lemma (in Group) Nset_Suc_im:"∀l ≤ (Suc n). G » (f l) ==>
                 insert (f (Suc n)) (f ` {i. i ≤ n}) = f ` {i. i ≤ (Suc n)}"
apply (rule equalityI)
 apply (rule subsetI)
  apply (simp add:image_def)
  apply (erule disjE) apply blast
  apply (erule exE, erule conjE)
  apply (frule_tac x = xa and y = n and z = "Suc n" in le_less_trans,
         simp,
         frule_tac x = xa and y = "Suc n" in less_imp_le, blast)
 apply (cut_tac Nset_Suc [of "n"], simp)
done

constdefs
  NfuncPair_neq_at::"[nat => 'a set, nat => 'a set, nat] => bool"
    "NfuncPair_neq_at f g i == (f i ≠  g i)" 

lemma LeastTr0:"[| (i::nat) < (LEAST l. P (l))|] ==> ¬ P (i)"
apply (rule not_less_Least)
apply simp
done

lemma (in Group) funeq_LeastTr1:"[|∀l≤ n. G » f l; ∀l ≤ n. G » g l; 
  (l :: nat) < (LEAST k. (NfuncPair_neq_at f g k)) |] ==> f l = g l"
apply (rule contrapos_pp, simp+)
apply (frule  LeastTr0 [of "l" "NfuncPair_neq_at f g"])
apply (simp add:NfuncPair_neq_at_def)
done

lemma (in Group) funeq_LeastTr1_1:"[|∀l ≤ (n::nat). G » f l; ∀l ≤ n. G »  g l; 
  (l :: nat) < (LEAST k. (f k ≠  g k)) |] ==> f l = g l"
apply (rule funeq_LeastTr1[of "n" "f" "g" "l"], assumption+)
apply (simp add:NfuncPair_neq_at_def)
done

lemma (in Group) Nfunc_LeastTr2_1:"[|i ≤ n; ∀l ≤ n. G » f l; ∀l ≤ n. G » g l;
       NfuncPair_neq_at f g i|] ==> 
        NfuncPair_neq_at f g (LEAST k. (NfuncPair_neq_at f g k))" 
apply (simp add: LeastI [of "NfuncPair_neq_at f g" "i"])
done

lemma (in Group) Nfunc_LeastTr2_2:"[|i ≤ n; ∀l ≤ n. G » f l; ∀l ≤ n. G » g l;
        NfuncPair_neq_at f g i|] ==> 
                          (LEAST k. (NfuncPair_neq_at f g k)) ≤ i" 
apply (simp add: Least_le [of "NfuncPair_neq_at f g" "i"])
done

lemma (in Group) Nfunc_LeastTr2_2_1:"[|i ≤ (n::nat); ∀l ≤ n. G » f l;
       ∀l ≤ n. G » g l; f i ≠ g i|] ==> (LEAST k. (f k ≠ g k)) ≤ i"
apply (rule contrapos_pp, simp+)
apply (simp add:nat_not_le_less)
apply (frule Nfunc_LeastTr2_2 [of "i" "n" "f" "g"], assumption+)
apply (simp add:NfuncPair_neq_at_def)+
done

lemma (in Group) Nfunc_LeastTr2_3:"[|∀l ≤ (n::nat). G » f l; ∀l ≤ n. G » g l; 
      i ≤ n; f i ≠ g i|]  ==>  
      f (LEAST k. (f k ≠  g k)) ≠ g (LEAST k. (f k ≠  g k))" 
apply (frule  Nfunc_LeastTr2_1 [of "i" "n" "f" "g"], assumption+)
apply (simp add:NfuncPair_neq_at_def)+
done

lemma  (in Group) Nfunc_LeastTr2_4:"[|∀l ≤ (n::nat). G » f l; ∀l ≤ n. G » g l; 
      i ≤ n; f i ≠ g i|] ==>(LEAST k. (f k ≠ g k)) ≤ n" 
apply (frule_tac i = i in Nfunc_LeastTr2_2 [of  _ "n" "f" "g"], 
                                      assumption+)
apply (simp add:NfuncPair_neq_at_def)
apply (frule le_trans [of "(LEAST k. NfuncPair_neq_at f g k)" "i" "n"],
                                  assumption+)
apply (simp add:NfuncPair_neq_at_def)
done

lemma (in Group) Nfunc_LeastTr2_5:"[|∀l≤ (n::nat). G » f l; ∀l ≤ n. G » g l; 
      ∃i ≤ n. (f i ≠ g i)|]  ==>  
      f (LEAST k. (f k ≠ g k)) ≠ g ((LEAST k. f k ≠ g k))"
apply (erule exE)
apply (rule_tac i = i in Nfunc_LeastTr2_3[of n f g], assumption+, simp+)
done

lemma (in Group) Nfunc_LeastTr2_6:"[|∀l ≤ (n::nat). G » f l; ∀l ≤ n. G » g l;
       ∃i ≤ n. (f i ≠ g i)|] ==> (LEAST k. (f k ≠ g k)) ≤ n"
apply (erule exE)
apply (rule_tac i = i in  Nfunc_LeastTr2_4, assumption+, simp+)
done

lemma (in Group) Nfunc_Least_sym:"[|∀l ≤ (n::nat). G » f l; ∀l ≤ n. G » g l; 
     ∃i ≤ n. (f i ≠ g i)|] ==> 
           (LEAST k. (f k ≠ g k)) = (LEAST k. (g k ≠ f k))"
apply (erule exE)
 apply (frule_tac i = i in Nfunc_LeastTr2_4 [of n f g], assumption+,
        simp+,
        frule_tac i = i in Nfunc_LeastTr2_3 [of n f g _], assumption+,
        simp+,
        frule_tac i = i in Nfunc_LeastTr2_4 [of n g f], assumption+,
        simp+, rule not_sym, simp,
        frule_tac i = i in Nfunc_LeastTr2_3 [of n g f _], assumption+,
        simp+, rule not_sym, simp)
 apply (frule_tac i = "(LEAST k. f k ≠ g k)" in 
           Nfunc_LeastTr2_2_1 [of _ "n" "g" "f"], assumption+,
         rule not_sym, simp) apply (
        frule_tac i = "(LEAST k. g k ≠ f k)" in 
           Nfunc_LeastTr2_2_1 [of _ n f g], assumption+,
           rule not_sym, simp)  
 apply (rule le_anti_sym, assumption+)
done

lemma Nfunc_iNJTr:"[|inj_on g {i. i ≤ (n::nat)}; i ≤ n; j ≤ n; i < j |] ==> g i ≠ g j"
apply (unfold inj_on_def)
apply (simp add:CollectI)
apply (rule contrapos_pp, simp+)
apply (frule_tac a = i in forall_spec) 
apply (frule_tac x = i and y = j and z = n in less_le_trans, assumption+,
       simp add:less_imp_le,
       thin_tac "∀x≤n. ∀y≤n. g x = g y --> x = y",
       rotate_tac -1,
       frule_tac a = j in forall_spec, assumption,
       thin_tac "∀y≤n. g i = g y --> i = y")
apply simp
done

lemma (in Group) Nfunc_LeastTr2_7:"[|∀l ≤ (n::nat). G » f l; ∀l ≤ n. G » g l; 
      inj_on g {i. i ≤ n}; ∃i ≤ n. (f i ≠ g i); 
      f k = g (LEAST k.(f k ≠ g k))|] ==>(LEAST k.(f k ≠ g k)) < k"
apply (rule contrapos_pp, simp+) 
apply (simp add:nat_not_le_less[THEN sym, of "LEAST k. f k ≠ g k" "k"])
apply (frule le_imp_less_or_eq)
apply (case_tac "k = (LEAST k. f k ≠ g k)")
 apply simp 
 apply (frule  Nfunc_LeastTr2_5 [of "n" "f" "g"], assumption+)
 apply simp 
apply (frule funeq_LeastTr1_1 [of "n" "f" "g" "k"], assumption+)
 apply simp
apply (frule Nfunc_LeastTr2_6 [of "n" "f" "g"], assumption+)
 apply simp
 apply (frule_tac le_trans[of "k" "LEAST k. f k ≠ g k" "n"], assumption+)
 apply (frule mem_of_Nset[of "k" "n"])
 apply (simp add:inj_on_def[of "g"])
done

lemma (in Group) Nfunc_LeastTr2_8:"[|∀l ≤ n. G » f l; ∀l ≤ n. G » g l; 
     inj_on g {i. i ≤ n}; ∃i ≤ n. f i ≠ g i; f `{i. i ≤ n} = g `{i. i ≤ n}|]
 ==> 
  ∃ k ∈(nset (Suc (LEAST i. (f i ≠ g i))) n). f k = g (LEAST i. (f i ≠ g i))"
apply (frule_tac Nfunc_LeastTr2_6 [of "n" "f" "g"], assumption+)
apply (cut_tac mem_in_image2[of "LEAST k. f k ≠ g k" "{i. i ≤ n}" "g"])
 apply (frule sym, thin_tac "f ` {i. i ≤ n} = g ` {i. i ≤ n}", simp)
 apply (thin_tac "g ` {i. i ≤ n} = f ` {i. i ≤ n}")
 apply (simp add:image_def)
 apply (rotate_tac -1, erule exE)
 apply (frule_tac k = x in Nfunc_LeastTr2_7[of "n" "f" "g"], assumption+)
 apply (erule conjE) apply (rule sym, assumption)
 apply (frule_tac m = "(LEAST k. f k ≠ g k)" and n = x in Suc_leI)
 apply (simp add:nset_def)
apply blast apply simp
done

lemma (in Group) ex_redchainTr1:"[|d_gchain G n f; 
       D_gchain G (card (f ` {i. i ≤ n}) - Suc 0) g; 
       g ` {i. i ≤ (card (f ` {i. i ≤ n}) - Suc 0)} = f ` {i. i ≤ n}|] ==> 
       g (card (f ` {i. i ≤ n}) - Suc 0) = f n" 
apply (case_tac "n = 0", simp, simp)
apply (rule contrapos_pp, simp+)
apply (cut_tac n_in_Nsetn[of "card (f ` {i. i ≤ n}) - Suc 0"])
apply (frule mem_in_image2[of "card (f ` {i. i ≤ n}) - Suc 0" 
                           "{i. i ≤ (card (f ` {i. i ≤ n}) - Suc 0)}" "g"])
 apply (cut_tac n_in_Nsetn[of "n"])
 apply (frule mem_in_image2[of "n" "{i. i ≤ n}" "f"])
 apply simp
 apply (simp add:image_def[of "f" "{i. i ≤ n}"])
 apply (erule exE)
 apply (frule_tac l = x in d_gchainTr2[of "n" "f" _ "n"], assumption+)
 apply simp+
 apply (erule conjE)
 apply (rotate_tac -1, frule sym,
        thin_tac "g (card {y. ∃x≤n. y = f x} - Suc 0) = f x",
        simp,
        thin_tac "f x = g (card {y. ∃x≤n. y = f x} - Suc 0)")
 apply (cut_tac mem_in_image2[of "n" "{i. i ≤ n}" "f"],
        unfold image_def)
 apply (frule sym,
        thin_tac "{y. ∃x∈{i. i ≤ card {y. ∃x≤n. y = f x} - Suc 0}. y = g x} =
         {y. ∃x≤n. y = f x}")
 apply (cut_tac eq_set_inc[of "f n" "{y. ∃x ≤ n. y = f x}"
         "{y. ∃x∈{i. i ≤ card {y. ∃x≤n. y = f x} - Suc 0}. y = g x}"]) 
 apply (thin_tac "{y. ∃x≤n. y = f x} =
         {y. ∃x∈{i. i ≤ card {y. ∃x≤n. y = f x} - Suc 0}. y = g x}",
        thin_tac "f n ∈ {y. ∃x∈{i. i ≤ n}. y = f x}")
apply  (simp, erule exE, erule conjE)
 apply (case_tac "xa ≤ card {y. ∃x≤n. y = f x} - Suc 0", simp)
 apply (frule D_gchain_is_d_gchain[of "card {y. ∃x≤n. y = f x} - Suc 0" g])
 apply (case_tac "card {y. ∃x ≤ n. y = f x} - Suc 0 = 0", 
         simp)
 apply (frule nat_nonzero_pos[of "card {y. ∃x ≤ n. y = f x} - Suc 0"],
        thin_tac "card {y. ∃x≤n. y = f x} - Suc 0 ≠ 0")
 apply (frule_tac l = xa in d_gchainTr2[of 
  "card {y. ∃x ≤ n. y = f x} - Suc 0" "g" _ 
                 "card {y. ∃x ≤ n. y = f x} - Suc 0"], assumption+)
 apply simp apply simp
 apply (frule_tac A = "g xa" and 
        B = "g (card {y. ∃x ≤ n. y = f x} - Suc 0)" in equalityI,
        assumption+, simp) apply simp+
done

lemma (in Group) ex_redchainTr1_1:"[|d_gchain G (n::nat) f; 
       D_gchain G (card (f ` {i. i ≤ n}) - Suc 0) g; 
       g ` {i. i ≤ (card (f ` {i. i ≤ n}) - Suc 0)} = f ` {i. i ≤ n}|] ==>
       g 0 = f 0"
apply (cut_tac Nset_inc_0[of "n"],
        frule mem_in_image2[of "0" "{i. i ≤ n}" "f"]) apply (
        frule sym) apply ( 
       thin_tac "g ` {i. i ≤ card (f ` {i. i ≤ n}) - Suc 0} = f ` {i. i ≤ n}")
  apply (
       frule eq_set_inc[of "f 0" "f ` {i. i ≤ n}" 
                 "g ` {i. i ≤ card (f ` {i. i ≤ n}) - Suc 0}"], assumption)
apply (       
 thin_tac "f 0 ∈ f ` {i. i ≤ n}",
        thin_tac "0 ∈ {i. i ≤ n}")
apply (cut_tac Nset_inc_0[of "card (f ` {i. i ≤ n}) - Suc 0"],
        frule mem_in_image2[of "0" "{i. i ≤ (card (f ` {i. i ≤ n}) - Suc 0)}"
         "g"],
        frule sym) apply (
       frule eq_set_inc[of "g 0" "g ` {i. i ≤ card (f ` {i. i ≤ n}) - Suc 0}" 
               "f ` {i. i ≤ n}"], assumption) apply (
       thin_tac "f ` {i. i ≤ n} = g ` {i. i ≤ card (f ` {i. i ≤ n}) - Suc 0}")
 apply (
       thin_tac "0 ∈ {i. i ≤ card (f ` {i. i ≤ n}) - Suc 0}") apply (
       thin_tac "g ` {i. i ≤ card (f ` {i. i ≤ n}) - Suc 0} = f ` {i. i ≤ n}")
 apply (
       thin_tac "g 0 ∈ g ` {i. i ≤ card (f ` {i. i ≤ n}) - Suc 0}")
apply (case_tac "n = 0", simp)
 apply (simp)
 apply (cut_tac mem_in_image3[of "f 0" "g" 
                "{i. i ≤ card (f ` {i. i ≤ n}) - Suc 0}"],
        frule mem_in_image3[of "g 0" "f" 
                    "{i. i ≤ n}"]) apply (
        thin_tac "f 0 ∈ g ` {i. i ≤ card (f ` {i. i ≤ n}) - Suc 0}",
        thin_tac "g 0 ∈ f ` {i. i ≤ n}") apply (erule bexE)+ apply (
        frule_tac j = aa in d_gchainTr2[of "n" "f" "0"], assumption+) 
  apply simp+ 
apply (rotate_tac -2, frule sym, thin_tac "g 0 = f aa", simp)
apply (case_tac "a = 0", simp)
apply (simp,
       frule_tac j = a in D_gchain0[of "card (f ` {i. i ≤ n}) - Suc 0" g 0],
       simp add:Nset_inc_0, assumption+,
       simp add:psubset_contr) 
apply simp
done

lemma (in Group) ex_redchainTr2:"d_gchain G (Suc n) f 
               ==> D_gchain G 0 (constmap {0::nat} {f (Suc n)})"
apply (simp add:D_gchain_def constmap_def)
apply (simp add:d_gchain_def)
done

lemma (in Group) last_mem_excluded:"[|d_gchain G (Suc n) f; f n ≠ f (Suc n)|] ==>
                                   f (Suc n) ∉ f ` {i. i ≤ n}"
apply (rule contrapos_pp, simp+)
apply (frule mem_in_image3[of "f (Suc n)" "f" "{i. i ≤ n}"], erule bexE)
apply (cut_tac zero_less_Suc[of "n"]) 
apply (frule_tac l = a in d_gchainTr2[of "Suc n" "f" _ "n"], assumption+)
 apply simp+ 
 apply (frule sym, thin_tac "f (Suc n) = f a", simp) 
apply (cut_tac l = n and j = "Suc n" in  d_gchainTr2[of "Suc n" "f"])
 apply simp+
done

lemma (in Group) ex_redchainTr4:"[|d_gchain G (Suc n) f; f n ≠ f (Suc n)|] ==>
                card (f ` {i. i ≤ (Suc n)}) = Suc (card (f ` {i. i ≤ n}))"
apply (cut_tac image_Nset_Suc [of "f" "n"])
apply simp
apply (rule card_insert_disjoint)
apply (simp add:finite_Nset)
apply (simp add:last_mem_excluded)
done

lemma (in Group) ex_redchainTr5:"d_gchain G n f ==> 0 < card (f ` {i. i≤ n})"
apply (simp add:image_Nsetn_card_pos)
done

lemma (in Group) ex_redchainTr6:"∀f. d_gchain G n f --> 
      (∃g. D_gchain G (card (f `{i. i ≤ n}) - 1) g ∧ 
                  (g ` {i. i ≤ (card (f `{i. i ≤ n}) - 1)} = f `{i. i ≤ n}))"
apply (induct_tac n)
apply (rule allI, rule impI)
 apply (simp add:image_def)
   apply (simp add:D_gchain_def d_gchain_def)
   apply blast
(** n **)
apply (rule allI, rule impI)
apply (case_tac "f (Suc n) = f n")
 apply (cut_tac n = n in Nset_Suc)
 apply (cut_tac n = n in n_in_Nsetn,
        frule_tac a = n and A = "{i. i ≤ n}" and f = f in mem_in_image2)
 apply (frule sym) apply (thin_tac "f (Suc n) = f n", simp)
 apply (subst insert_absorb, assumption)+
 apply (frule_tac n = n and f = f in d_gchain_pre, blast)

apply (frule_tac n = n and f = f in d_gchain_pre)
 apply (frule_tac a = f in forall_spec, assumption,
        thin_tac "∀f. d_gchain G n f -->
               (∃g. D_gchain G (card (f ` {i. i ≤ n}) - 1) g ∧
                    g ` {i. i ≤ card (f ` {i. i ≤ n}) - 1} =
                    f ` {i. i ≤ n})")
 apply (erule exE, erule conjE)
apply (simp add:image_Nset_Suc)

apply (frule_tac n = n and f = f in ex_redchainTr2)
apply (frule_tac n = "card (f ` {i. i ≤ n}) - Suc 0" and f = g and 
       g = "constmap {0} {f (Suc n)}" in joint_D_gchain_n0, assumption+)
apply (simp add: ex_redchainTr1)
 apply (simp add: constmap_def Nset_inc_0) 
 apply (cut_tac n = n in zero_less_Suc) 
 apply (frule_tac n = "Suc n" and f = f and l = n and j = "Suc n" in 
        d_gchainTr2, assumption+)
        apply simp  apply simp  apply simp
 apply (simp add:psubset_eq)
 apply (cut_tac f = f and n = n in image_Nsetn_card_pos,
        cut_tac n = n in finite_Nset, 
        frule_tac F = "{i. i ≤ n}" and h = f in finite_imageI,
        frule_tac n = n and f = f in last_mem_excluded,
        rule not_sym, assumption)
 apply simp+
 apply (cut_tac n = "card (f ` {i. i ≤ n}) - Suc 0" and f = g and m = 0 and 
        g = "constmap {0} {f (Suc n)}" in im_jointfun1)
 apply simp
 apply (simp add:Nset_0 constmap_def)
apply blast
done

lemma (in Group) ex_redchain:"d_gchain G n f ==>
      (∃g. D_gchain G (card (f ` {i. i ≤ n}) - 1) g ∧ 
       g ` {i. i ≤ (card (f ` {i. i ≤ n}) - 1)} = f ` {i. i ≤ n})"
apply (cut_tac ex_redchainTr6 [of "n"])
apply simp
done

lemma (in Group) const_W_cmpser:"d_gchain G (Suc n) f ==>
                         W_cmpser G 0 (constmap {0::nat} {f (Suc n)})"
apply (simp add:W_cmpser_def d_gchain_def constmap_def)
done

lemma (in Group) ex_W_cmpserTr0m:"∀f. w_cmpser G m f -->  
  (∃g. (W_cmpser G (card (f `{i. i ≤ m}) - 1) g ∧ 
           g `{i. i ≤ (card (f `{i. i ≤ m}) - 1)} = f `{i. i ≤ m}))"
apply (induct_tac m)
(********* induct_step m = 0 ***************)
apply (rule allI, rule impI) 
 apply simp 
 apply (simp add:w_cmpser_def W_cmpser_def)
 apply blast
(********** induct_step  m ****************)
apply (rule allI, rule impI)  
 apply (case_tac "f (Suc n) = f n")
 apply (cut_tac n = n in Nset_Suc)
 apply (cut_tac n = n in n_in_Nsetn,
        frule_tac a = n and A = "{i. i ≤ n}" and f = f in mem_in_image2)
 apply (frule sym) apply (thin_tac "f (Suc n) = f n", simp)
 apply (subst insert_absorb, assumption)+
 apply (frule_tac n = n and f = f in  w_cmpser_pre, blast)

apply (frule_tac n = n and f = f in w_cmpser_pre)
 apply (frule_tac a = f in forall_spec, assumption,
        thin_tac "∀f. w_cmpser G n f -->
               (∃g. W_cmpser G (card (f ` {i. i ≤ n}) - 1) g ∧
                    g ` {i. i ≤ card (f ` {i. i ≤ n}) - 1} =
                    f ` {i. i ≤ n})")
 apply (erule exE, erule conjE)
apply (simp add:image_Nset_Suc,
       frule_tac n = "Suc n" and f = f in w_cmpser_is_d_gchain,
       frule_tac n = n and f = f in const_W_cmpser)
apply (frule_tac n = "card (f ` {i. i ≤ n}) - Suc 0" and f = g and 
       g = "constmap {0::nat} {f (Suc n)}" in joint_W_cmpser_n0, assumption+)
 apply (frule_tac n = "card (f ` {i. i ≤ n}) - Suc 0" and f = g in 
                      W_cmpser_is_D_gchain)
 apply (frule d_gchain_pre)
 apply (subst ex_redchainTr1, assumption+)
 apply (simp add:constmap_def Nset_inc_0)
 apply (simp add:w_cmpser_def)
 apply (frule d_gchain_pre)
 apply (frule_tac n = "card (f ` {i. i ≤ n}) - Suc 0" and f = g in 
                      W_cmpser_is_D_gchain)
 apply (frule_tac n = n and f = f and g = g in ex_redchainTr1, assumption+)
        apply simp
 apply (simp add:constmap_def Nset_inc_0,
        thin_tac "d_gchain G n f", simp add:d_gchain_def) 
 apply (cut_tac n = "Suc n" in n_in_Nsetn,
        frule_tac a = "Suc n" in forall_spec1, simp,
        simp add:psubset_eq)
 apply (cut_tac f = f and n = n in image_Nsetn_card_pos,
        cut_tac n = n in finite_Nset, 
        frule_tac F = "{i. i ≤ n}" and h = f in finite_imageI,
        frule_tac n = n and f = f in last_mem_excluded,
        rule not_sym, assumption)
 apply simp+
 apply (cut_tac n = "card (f ` {i. i ≤ n}) - Suc 0" and f = g and m = 0 and 
        g = "constmap {0::nat} {f (Suc n)}" in im_jointfun1)
 apply simp
 apply (simp add:Nset_0 constmap_def)
apply blast
done

lemma (in Group) ex_W_cmpser:"w_cmpser G m f ==>
       ∃g. W_cmpser G (card (f ` {i. i ≤ m}) - 1) g  ∧  
              g ` {i. i ≤ (card (f ` {i. i ≤ m}) - 1)} = f ` {i. i ≤ m}"
apply (cut_tac ex_W_cmpserTr0m [of "m"])
apply simp
done

section "17. Existence of reduced chain and composition series"

lemma (in Group) ex_W_cmpserTr3m1:"[|tw_cmpser G (m::nat) f; 
       W_cmpser G ((card (f ` {i. i ≤ m})) - 1) g; 
       g ` {i. i ≤ ((card (f ` {i. i ≤ m})) - 1)} = f `{i. i ≤ m}|] ==> 
      tW_cmpser G ((card (f ` {i. i ≤ m})) - 1) g"
apply (frule_tac tw_cmpser_is_w_cmpser [of "m" "f"],
       frule_tac w_cmpser_is_d_gchain [of "m" "f"],
       frule_tac W_cmpser_is_D_gchain [of "(card (f ` {i. i ≤ m}) - 1)" "g"])

apply (frule ex_redchainTr1 [of "m" "f" "g"])
 apply simp+
apply (frule_tac ex_redchainTr1_1 [of "m" "f" "g"])
 apply (simp add:tW_cmpser_def tw_cmpser_def)
 apply (case_tac "m = 0") apply simp
 apply (insert finite_Nset [of "0"])
 apply (cut_tac card_image_le [of "{0::nat}" "f"])
 apply (simp add:card_Nset [of "0"], simp) 
apply (simp add:tW_cmpser_def)
 apply (case_tac "card (f ` {i. i ≤ m}) ≤ Suc 0") apply simp
 apply (simp add:td_gchain_def tw_cmpser_def) 
 apply (case_tac "m = 0") 
 apply (thin_tac "f 0 = f m", thin_tac "g 0 = f m") apply simp
 apply ( simp add:td_gchain_def) apply ( erule conjE, simp)
 apply simp
 apply (simp add:td_gchain_def[of "G" "m" "f"], (erule conjE)+, simp)
apply simp
 apply (simp add:tD_gchain_def tw_cmpser_def td_gchain_def [of _ "m" "f"])
apply (case_tac "m = 0", simp add:card1)
 apply (simp, erule conjE, simp add:td_gchain_def)
 apply (simp add:W_cmpser_def)
done

lemma (in Group) ex_W_cmpserTr3m:"tw_cmpser G m f ==> 
           ∃g. tW_cmpser G ((card (f ` {i. i ≤ m})) - 1) g ∧ 
               g `{ i. i ≤ (card (f `{i. i ≤ m}) - 1)} = f ` {i. i ≤ m}"
apply (frule tw_cmpser_is_w_cmpser[of "m" "f"])
apply (frule ex_W_cmpser [of "m" "f"])
apply (auto del:equalityI)
apply (frule_tac g = g in ex_W_cmpserTr3m1 [of "m" "f"])
apply simp+ apply blast
done

constdefs (structure G) 
  red_ch_cd::"[_ , nat => 'a set, nat, nat => 'a set ] => bool"
    "red_ch_cd G f m g == tW_cmpser G (card (f ` {i. i ≤ m}) - 1) g ∧ 
                 (g `{i . i ≤ (card (f ` {i. i ≤ m}) - 1)} = f` {i. i ≤ m})"
 
  red_chain::"[_ , nat, nat => 'a set] => (nat => 'a set)"
  "red_chain G m f == (SOME g. g ∈ {h. red_ch_cd G f m h})"

lemma (in Group) red_chainTr0m1_1:"tw_cmpser G m f ==> 
       (SOME g. g ∈ {h. red_ch_cd G f m h}) ∈ {h. red_ch_cd G f m h}"
apply (rule nonempty_some [of "{h. red_ch_cd G f m h}"])
apply (frule ex_W_cmpserTr3m [of "m" "f"])
 apply simp
 apply (simp add:red_ch_cd_def)
done

lemma (in Group) red_chain_m:"tw_cmpser G m f ==> 
      tW_cmpser G (card (f ` {i. i ≤ m}) - 1) (red_chain G m f) ∧ 
      (red_chain G m f) `{i. i ≤ (card (f `{i. i ≤ m}) - 1)} = f ` {i. i ≤ m}"
apply (frule red_chainTr0m1_1 [of "m" "f"])
apply simp
apply (simp add:red_ch_cd_def)
apply (simp add:red_chain_def)
done

section "18. Chain of groups II"

constdefs  
 Gchain :: "[nat, nat => (('a set), 'more) Group_scheme] => bool"
 "Gchain n g == ∀l ≤ n. Group (g l)"  

 isom_Gchains::"[nat, nat => nat, nat => (('a set), 'more) Group_scheme, 
            nat => (('a set), 'more) Group_scheme]  => bool"
 "isom_Gchains n f g h == ∀i ≤ n. (g i) ≅ (h (f i))"
(* g, h are sequences of groups and f is gbijection of Nset to Nset *)

 Gch_bridge::"[nat, nat => (('a set), 'more) Group_scheme, nat => 
             (('a set), 'more) Group_scheme, nat => nat]  => bool"
 "Gch_bridge n g h f == (∀l ≤ n. f l ≤ n) ∧ inj_on f {i. i ≤ n} ∧ 
                         isom_Gchains n f g h" 

lemma Gchain_pre:"Gchain (Suc n) g ==> Gchain n g"    
apply (simp add:Gchain_def)
done

lemma (in Group) isom_unit:"[|G » H; G » K; H = {\<one>}|]  ==> 
                                    Gp G H ≅ Gp G K --> K = {\<one>}"
apply (simp add:isomorphic_def)
apply (rule impI)
apply (erule exE)
 apply (simp add:gbijec_def)
 apply (erule conjE)
 apply (simp add:gsurjec_def ginjec_def)
 apply (erule conjE)
apply (simp add:Gp_carrier)
apply (simp add:surj_to_def)
apply (cut_tac a = "f \<one>" in finite1)

apply (frule sg_unit_closed [of "K"])
 apply (frule singleton_sub[of "\<one>" "K"])
 apply (rotate_tac 4, frule sym, thin_tac "{f \<one>} = K") 
 apply (rule card_seteq[THEN sym, of "K" "{\<one>}"])
 apply (simp add:finite1) apply assumption
 apply (simp add:card1)
done

lemma isom_gch_unitsTr4:"[|Group F; Group G; Ugp E; F ≅ G; F ≅ E|] ==>
                                 G ≅ E"
apply (simp add:Ugp_def)
apply (erule conjE)
apply (frule isomTr1 [of "F" "G"], assumption+)
apply (rule isomTr2 [of "G" "F" "E"], assumption+)
done 

lemma isom_gch_cmp:"[|Gchain n g; Gchain n h; f1 ∈ {i. i ≤ n} -> {i. i ≤ n}; 
           f2 ∈ {i. i ≤ n} -> {i. i ≤ n}; isom_Gchains n (cmp f2 f1) g h|] ==> 
             isom_Gchains n f1 g (cmp h f2)"
apply (simp add:isom_Gchains_def)
apply (simp add:cmp_def)
done

lemma isom_gch_transp:"[|Gchain n f; i ≤ n; j ≤ n; i < j|] ==> 
                 isom_Gchains n (transpos i j) f (cmp  f (transpos i j))"
apply (rule isom_gch_cmp [of "n" "f" _ "transpos i j" "transpos i j"],
                                                                 assumption+)
 apply (rule transpos_hom, assumption+) apply simp
 apply (rule transpos_hom, assumption+) apply simp
apply (simp add:isom_Gchains_def)
apply (rule allI, rule impI) 
apply (frule less_le_trans [of "i" "j" "n"], assumption+)
apply (frule less_imp_le [of "i" "n"])
apply (frule_tac k = ia in cmp_transpos1 [of "i" "n" "j"], assumption+)
 apply simp+ 
apply (simp add:Gchain_def)
done

lemma isom_gch_units_transpTr0:"[|Ugp E; Gchain n g; Gchain n h; i ≤ n; j ≤ n;
 i < j; isom_Gchains n (transpos i j) g h|] ==>
 {i. i ≤ n ∧ g i ≅ E} - {i, j} ={i. i ≤ n ∧ h i ≅ E} - {i, j}"
apply (simp add:isom_Gchains_def)
apply (rule equalityI)
 apply (rule subsetI, simp add:CollectI)
 apply (erule conjE)+
 apply (cut_tac x = x in transpos_id_1 [of "i" "n" "j"], simp+)
apply (subgoal_tac "g x ≅ h (transpos i j x)",
       thin_tac "∀ia≤n. g ia ≅ h (transpos i j ia)", simp)
apply (subgoal_tac "Group (g x)", subgoal_tac "Group (h x)")
apply (simp add:Ugp_def) apply (erule conjE)
apply (frule_tac F = "g x" and G = "h x" in isomTr1, assumption+)
apply (rule_tac F = "h x" and G = "g x" and H = E in isomTr2, assumption+)
apply (simp add:Gchain_def, simp add:Gchain_def)
 apply (thin_tac "transpos i j x = x")
 apply simp
apply (rule subsetI, simp add:CollectI)
 apply (erule conjE)+
 apply (cut_tac x = x in transpos_id_1 [of "i" "n" "j"], simp+)
apply (subgoal_tac "g x ≅ h (transpos i j x)",
       thin_tac "∀ia≤n. g ia ≅ h (transpos i j ia)")
 apply simp
apply (subgoal_tac "Group (g x)",
       subgoal_tac "Group (h x)")
apply (simp add:Ugp_def) apply (erule conjE)
apply (rule_tac F = "g x" and G = "h x" and H = E in isomTr2, assumption+)
apply (simp add:Gchain_def, simp add:Gchain_def)
 apply (thin_tac "transpos i j x = x")
 apply simp
done

lemma isom_gch_units_transpTr1:"[|Ugp E; Gchain n g; i ≤ n; j ≤ n; g j ≅ E; 
      i ≠ j|] ==> 
      insert j ({i. i ≤ n ∧ g i ≅ E} - {i, j}) = {i. i ≤ n ∧ g i ≅ E} - {i}"
apply (rule equalityI)
 apply (rule subsetI) apply (simp add:CollectI)
 apply blast
apply (rule subsetI) apply (simp add:CollectI)
done

lemma isom_gch_units_transpTr2:"[|Ugp E; Gchain n g; i ≤ n; j ≤ n; i < j;
      g i ≅ E|] ==> 
      {i. i ≤ n ∧ g i ≅ E} = insert i ({i. i ≤ n ∧  g i ≅ E} - {i})"
apply (rule equalityI)
 apply (rule subsetI, simp add:CollectI)
 apply (rule subsetI, simp add:CollectI)
 apply (erule disjE, simp)
 apply simp
done

lemma isom_gch_units_transpTr3:"[|Ugp E; Gchain n g; i ≤ n|]
                         ==> finite ({i. i ≤ n ∧ g i ≅ E} - {i})"
apply (insert finite_Nset [of "n"])
apply (rule finite_subset[of "{i. i ≤ n ∧ g i ≅ E} - {i}" "{i. i ≤ n}"])
apply (rule subsetI, simp)
apply assumption
done  

lemma isom_gch_units_transpTr4:"[|Ugp E; Gchain n g; i ≤ n|]
                         ==> finite ({i. i ≤ n ∧ g i ≅ E} - {i, j})"
apply (insert finite_Nset [of "n"])
apply (rule finite_subset[of "{i. i ≤ n ∧ g i ≅ E} - {i, j}" "{i. i ≤ n}"])
apply (rule subsetI, simp)
apply assumption
done

lemma isom_gch_units_transpTr5_1:"[|Ugp E; Gchain n g; Gchain n h; i ≤ (n::nat);
      j ≤ n; i < j; isom_Gchains n (transpos i j) g h|] ==> g i ≅ h j"
apply (simp add:isom_Gchains_def)
apply (frule_tac a = i in forall_spec,
       frule_tac x = i and y = j and z = n in less_le_trans,
       assumption+, simp,
       thin_tac "∀ia ≤ n. g ia ≅ h (transpos i j ia)")
apply (simp add:transpos_ij_1 [of "i" "n" "j"])
done

lemma isom_gch_units_transpTr5_2:"[|Ugp E; Gchain n g; Gchain n h; i ≤ n; 
      j ≤ n; i < j; isom_Gchains n (transpos i j) g h|] ==> g j ≅ h i"
apply (simp add:isom_Gchains_def)
apply (frule_tac a = j in forall_spec1,
       thin_tac "∀ia≤ n. g ia ≅ h (transpos i j ia)")
apply (simp add:transpos_ij_2 [of "i" "n" "j"]) 
done

lemma isom_gch_units_transpTr6:"[|Gchain n g; i ≤ n|] ==> Group (g i)"
apply (simp add:Gchain_def)
done

lemma isom_gch_units_transpTr7:"[|Ugp E; i ≤ n; j ≤ n; g j ≅ h i; 
      Group (h i); Group (g j); ¬ g j ≅ E|] ==>  ¬ h i ≅ E"
apply (rule contrapos_pp, simp+)
apply (frule isomTr2 [of "g j" "h i" "E"], assumption+)
apply (simp add:Ugp_def)
apply assumption+
apply simp
done

lemma isom_gch_units_transpTr8_1:"[|Ugp E; Gchain n g; i ≤ n; j ≤ n; 
      g i ≅ E; ¬ g j ≅ E|] ==> 
      {i. i ≤ n ∧ g i ≅ E} = {i. i ≤ n ∧ g i ≅ E} - { j }"
by auto

lemma isom_gch_units_transpTr8_2:"[|Ugp E; Gchain n g; i ≤ n; j ≤ n;
       ¬ g i ≅ E; ¬ g j ≅ E|] ==> 
       {i. i ≤ n ∧ g i ≅ E} = {i. i ≤ n ∧ g i ≅ E} - {i, j }"
by auto

lemma isom_gch_units_transp:"[|Ugp E; Gchain n g; Gchain n h; i ≤ n; j ≤ n; 
      i < j; isom_Gchains n (transpos i j) g h|] ==>  
       card {i. i ≤ n ∧ g i ≅ E} = card {i. i ≤ n ∧ h i ≅ E}" 
apply (frule isom_gch_units_transpTr0 [of "E" "n" "g" "h" "i" "j"], 
        assumption+)
apply (frule isom_gch_units_transpTr6 [of "n" "g" "i"], assumption+)
apply (frule isom_gch_units_transpTr6 [of "n" "h" "i"], assumption+)
apply (frule isom_gch_units_transpTr6 [of "n" "g" "j"], assumption+)
apply (frule isom_gch_units_transpTr6 [of "n" "h" "j"], assumption+)
apply (unfold Ugp_def) apply (frule conjunct1) apply (fold Ugp_def)
apply (frule isom_gch_units_transpTr5_1 [of "E" "n" "g" "h" "i" "j"], 
                                                          assumption+)
apply (frule isom_gch_units_transpTr5_2 [of "E" "n" "g" "h" "i" "j"], 
                                                          assumption+)
apply (case_tac "g i ≅ E")
 apply (case_tac "g j ≅ E")  (** g i ≅ E and g j ≅ E **)
 apply (subst isom_gch_units_transpTr2 [of "E" "n" "g" "i" "j"], assumption+)
 apply (subst isom_gch_units_transpTr2 [of "E" "n" "h" "i" "j"], assumption+) 
 apply (rule isom_gch_unitsTr4 [of "g j" "h i" "E"], assumption+)
 apply (subst card_insert_disjoint) 
 apply (rule isom_gch_units_transpTr3, assumption+)
 apply simp
 apply (subst card_insert_disjoint)
 apply (rule isom_gch_units_transpTr3, assumption+)
 apply simp
 apply (subst isom_gch_units_transpTr1[THEN sym, of "E" "n" "g" "i" "j"], assumption+) apply simp
apply (subst isom_gch_units_transpTr1[THEN sym, of "E" "n" "h" "i" "j"], assumption+)
 apply (rule isom_gch_unitsTr4 [of "g i" "h j" "E"], assumption+)
apply simp 
apply simp
apply (subst isom_gch_units_transpTr8_1 [of "E" "n" "g" "i" "j"], assumption+)
apply (subst isom_gch_units_transpTr8_1 [of "E" "n" "h" "j" "i"], assumption+)
 apply (rule isom_gch_unitsTr4 [of "g i" "h j" "E"], assumption+)
 apply (rule isom_gch_units_transpTr7 [of "E" "i" "n" "j" "g" "h"], 
                                                        assumption+)
 apply (subst isom_gch_units_transpTr1 [THEN sym, of "E" "n" "g" "j" "i"], assumption+) apply simp
 apply (subst card_insert_disjoint)
 apply (rule isom_gch_units_transpTr4, assumption+)
 apply simp
 apply (subst isom_gch_units_transpTr1 [THEN sym, of "E" "n" "h" "i" "j"],
                                                                assumption+)
 apply (rule isom_gch_unitsTr4 [of "g i" "h j" "E"], assumption+)
 apply simp
 apply (subst card_insert_disjoint)
  apply (rule isom_gch_units_transpTr4, assumption+)
  apply simp
  apply (insert Nset_2 [of "j" "i"])
  apply simp
apply (case_tac "g j ≅ E")
apply (subst isom_gch_units_transpTr8_1 [of "E" "n" "g" "j" "i"], assumption+)
apply (subst isom_gch_units_transpTr8_1 [of "E" "n" "h" "i" "j"], assumption+)
 apply (rule isom_gch_unitsTr4 [of "g j" "h i" "E"], assumption+)
 apply (rule isom_gch_units_transpTr7 [of "E" "j" "n" "i" "g" "h"], assumption+)
 apply (subst isom_gch_units_transpTr1 [THEN sym, of "E" "n" "g" "i" "j"],
                                                              assumption+)
 apply simp
 apply (subst isom_gch_units_transpTr1 [THEN sym, of "E" "n" "h" "j" "i"],
                                                              assumption+)
 apply (rule isom_gch_unitsTr4 [of "g j" "h i" "E"], assumption+)
 apply simp apply simp
 apply (subst card_insert_disjoint)
  apply (rule isom_gch_units_transpTr4, assumption+)
  apply simp apply simp
 apply (subst card_insert_disjoint)
  apply (rule isom_gch_units_transpTr4, assumption+)
  apply simp  apply simp apply simp
 apply (subst isom_gch_units_transpTr8_2 [of "E" "n" "g" "i" "j"], assumption+)
 apply (subst isom_gch_units_transpTr8_2 [of "E" "n" "h" "i" "j"], assumption+)
 apply (rule isom_gch_units_transpTr7[of "E" "i" "n" "j" "g" "h"], assumption+)
 apply (rule isom_gch_units_transpTr7[of "E" "j" "n" "i" "g" "h"], assumption+)
apply simp
done

lemma TR_isom_gch_units:"[|Ugp E; Gchain n f; i ≤ n; j ≤ n; i < j|] ==>
      card {k. k ≤ n ∧ f k ≅ E} = card {k. k ≤ n ∧ 
      (cmp f (transpos i j)) k ≅ E}" 
apply (frule isom_gch_transp [of "n" "f" "i" "j"], assumption+)
(* apply (subgoal_tac "Gchain n (cmp f (transpos i j))") *)
 apply (rule isom_gch_units_transp [of "E" "n" "f" _ "i" "j"], assumption+)
 apply (simp add:Gchain_def)
 apply (rule allI, rule impI)
 apply (simp add:cmp_def)
apply (cut_tac l = l in transpos_mem [of "i" "n" "j"],
       frule_tac x = i and y = j and z = n in less_le_trans, assumption+,
       simp)
 apply simp+
done

lemma TR_isom_gch_units_1:"[|Ugp E; Gchain n f; i ≤ n; j ≤ n; i < j|]  ==>  
      card {k. k ≤ n ∧ f k ≅ E} = card {k. k ≤ n ∧ f (transpos i j k) ≅ E}" 
apply (frule TR_isom_gch_units [of "E" "n" "f" "i" "j"], assumption+)
 apply (simp add:cmp_def)
done

lemma isom_tgch_unitsTr0_1:"[|Ugp E; Gchain (Suc n) g; g (Suc n) ≅ E|] ==> 
      {i. i ≤ (Suc n) ∧ g i ≅ E} = insert (Suc n) {i. i ≤ n ∧ g i ≅ E}"
apply (rule equalityI)
apply (rule subsetI)
 apply (simp add:CollectI)
 apply (case_tac "x = Suc n") apply simp
 apply (erule conjE) apply simp 
apply (rule subsetI)
 apply (simp add:CollectI)
 apply (erule disjE) apply simp
apply simp
done

lemma isom_tgch_unitsTr0_2:"Ugp E  ==> finite ({i. i ≤ (n::nat) ∧ g i ≅ E})"
apply (insert finite_Nset [of "n"])
(*apply (subgoal_tac "{i. i ∈ Nset n ∧ g i ≅ E} ⊆ Nset n")*)
apply (rule finite_subset[of "{i. i ≤ n ∧ g i ≅ E}" "{i. i ≤ n}"])
apply (rule subsetI, simp)
apply assumption
done

lemma isom_tgch_unitsTr0_3:"[|Ugp E; Gchain (Suc n) g; ¬ g (Suc n) ≅ E|]
      ==> {i. i ≤ (Suc n) ∧ g i ≅ E} = {i. i ≤ n ∧ g i ≅ E}"
apply (rule equalityI)
apply (rule subsetI, simp add:CollectI)
apply (case_tac "x = Suc n", simp, erule conjE)
 apply arith
apply (rule subsetI, simp add:CollectI) 
done

lemma isom_tgch_unitsTr0:"[|Ugp E; 
      card {i. i ≤ n ∧ g i ≅ E} = card {i. i ≤ n ∧ h i ≅ E}; 
      Gchain (Suc n) g ∧ Gchain (Suc n) h ∧ Gch_bridge (Suc n) g h f; 
      f (Suc n) = Suc n|] ==> 
      card {i. i ≤ (Suc n) ∧ g i ≅ E} = 
                                card {i. i ≤ (Suc n) ∧ h i ≅ E}"
apply (erule conjE)+
apply (frule isom_gch_units_transpTr6 [of "Suc n" "g" "Suc n"])
apply (simp add:n_in_Nsetn)
apply (frule isom_gch_units_transpTr6 [of "Suc n" "h" "Suc n"])
apply (simp add:n_in_Nsetn)
apply (simp add:Gch_bridge_def isom_Gchains_def)
 apply (erule conjE)+
 apply (rotate_tac -1, 
        frule_tac a = "Suc n" in forall_spec,
        thin_tac "∀i≤Suc n. g i ≅ h (f i)", simp+,
        thin_tac "∀i≤Suc n. g i ≅ h (f i)")
apply (case_tac "g (Suc n) ≅ E")
 apply (subst isom_tgch_unitsTr0_1 [of "E" "n" "g"], assumption+)
 apply (subst isom_tgch_unitsTr0_1 [of "E" "n" "h"], assumption+) 
 apply (frule isom_gch_unitsTr4 [of "g (Suc n)" "h (Suc n)" "E"], assumption+)
apply (subst card_insert_disjoint) 
 apply (rule finite_subset[of "{i. i ≤ n ∧ g i ≅ E}" "{i. i ≤ n}"])
 apply (rule subsetI, simp) apply (simp add:finite_Nset)
 apply simp
apply (subst card_insert_disjoint) 
 apply (rule finite_subset[of "{i. i ≤ n ∧ h i ≅ E}" "{i. i ≤ n}"])
 apply (rule subsetI, simp) apply (simp add:finite_Nset) apply simp
 apply simp
 apply (cut_tac isom_gch_units_transpTr7[of E "Suc n" "Suc n" "Suc n" g h])
 apply (subgoal_tac "{i. i ≤ Suc n ∧ g i ≅ E} = {i. i ≤ n ∧ g i ≅ E}",
        subgoal_tac "{i. i ≤ Suc n ∧ h i ≅ E} = {i. i ≤ n ∧ h i ≅ E}",
        simp)
 apply (rule equalityI, rule subsetI, simp,
        erule conjE, case_tac "x = Suc n", simp,
        frule_tac x = x and y = "Suc n" in le_imp_less_or_eq,
        thin_tac "x ≤ Suc n", simp,
        rule subsetI, simp)
 apply (rule equalityI, rule subsetI, simp,
        erule conjE, case_tac "x = Suc n", simp,
        frule_tac m = x and n = "Suc n" in noteq_le_less, assumption+,
        simp,        
        rule subsetI, simp)    
 apply simp+
done

lemma isom_gch_unitsTr1_1:" [|Ugp E; Gchain (Suc n) g ∧  Gchain (Suc n) h 
       ∧ Gch_bridge (Suc n) g h f; f (Suc n) = Suc n|] ==> 
           Gchain n g ∧ Gchain n h ∧ Gch_bridge n g h f"
apply (erule conjE)+
apply (frule Gchain_pre [of "n" "g"])
apply (frule Gchain_pre [of "n" "h"])
apply simp
apply (simp add:Gch_bridge_def) apply (erule conjE)+
apply (rule conjI)
apply (rule Nset_injTr2, assumption+)
apply (rule conjI)
apply (rule Nset_injTr1, assumption+)
apply (simp add:isom_Gchains_def)
done

lemma isom_gch_unitsTr1_2:"[|Ugp E; f (Suc n) ≠ Suc n; inj_on f {i. i≤(Suc n)};
        ∀l ≤ (Suc n). f l ≤ (Suc n)|] ==> 
        (cmp (transpos (f (Suc n)) (Suc n)) f) (Suc n) = Suc n"
apply (simp add:cmp_def)
apply (cut_tac n_in_Nsetn[of "Suc n"], simp)
apply (frule_tac a = "Suc n" in forall_spec, 
      simp,
      thin_tac "∀l≤Suc n. f l ≤ Suc n") 
apply (rule transpos_ij_1, assumption+, simp+)
done

lemma isom_gch_unitsTr1_3:"[|Ugp E; f (Suc n) ≠ Suc n; 
     ∀l ≤ (Suc n). f l ≤ (Suc n); inj_on f {i. i ≤ (Suc n)}|] ==> 
      inj_on (cmp (transpos (f (Suc n)) (Suc n)) f) {i. i ≤ (Suc n)}"
apply (cut_tac n_in_Nsetn[of "Suc n"], simp)
apply (frule_tac a = "Suc n" in forall_spec, simp)
apply (frule transpos_hom [of "f (Suc n)" "Suc n" "Suc n"], simp+)
    thm transpos_inj
apply (cut_tac transpos_inj [of "f (Suc n)" "Suc n" "Suc n"])
apply (cut_tac cmp_inj_1 [of f "{i. i ≤ (Suc n)}" "{i. i ≤ (Suc n)}"
       "transpos (f (Suc n)) (Suc n)" "{i. i ≤ (Suc n)}"])
apply simp+
apply (rule univar_func_test, rule ballI, simp)
apply assumption+ apply simp+
done

lemma isom_gch_unitsTr1_4:"[|Ugp E; f (Suc n) ≠ Suc n; inj_on f {i. i≤(Suc n)};
                   ∀l ≤ (Suc n). f l ≤ (Suc n)|] ==> 
                    inj_on (cmp (transpos (f (Suc n)) (Suc n)) f) {i. i ≤ n}"
apply (frule isom_gch_unitsTr1_3 [of "E" "f" "n"], assumption+)
apply (frule isom_gch_unitsTr1_2 [of "E" "f" "n"], assumption+)
apply (rule Nset_injTr1 [of "n" "(cmp (transpos (f (Suc n)) (Suc n)) f)"])
apply (rule allI, rule impI)
apply (simp add:cmp_def)  
apply (cut_tac l = "f l" in transpos_mem[of "f (Suc n)" "Suc n" "Suc n"], 
       simp+)
done

lemma isom_gch_unitsTr1_5:"[|Ugp E; Gchain (Suc n) g ∧ Gchain (Suc n) h ∧ 
      Gch_bridge (Suc n) g h f; f (Suc n) ≠ Suc n |] ==> 
      Gchain n g ∧ Gchain n (cmp h (transpos (f (Suc n)) (Suc n))) ∧ 
      Gch_bridge n g (cmp h (transpos (f (Suc n)) (Suc n))) 
                               (cmp (transpos (f (Suc n)) (Suc n)) f)"
apply (erule conjE)+ 
apply (simp add:Gchain_pre [of "n" "g"])
apply (rule conjI)
apply (simp add:Gchain_def) apply (rule allI, rule impI)
apply (simp add:Gch_bridge_def) apply (frule conjunct1)
apply (fold Gch_bridge_def)
apply (cut_tac n_in_Nsetn[of "Suc n"])
apply (cut_tac l = l in transpos_mem [of "f (Suc n)" "Suc n" "Suc n"])
      apply simp+
apply (simp add:cmp_def)

apply (simp add:Gch_bridge_def)
apply (erule conjE)+
apply (rule conjI)
apply (cut_tac n_in_Nsetn[of "Suc n"])
apply (rule allI, rule impI, simp add:cmp_def)
apply (frule isom_gch_unitsTr1_2 [of "E" "f" "n"], assumption+)
apply (frule isom_gch_unitsTr1_3 [of "E" "f" "n"], assumption+)
apply (cut_tac l = "f l" in transpos_mem[of "f (Suc n)" "Suc n" "Suc n"])
 apply simp+ 
 apply (simp add:inj_on_def[of "cmp (transpos (f (Suc n)) (Suc n)) f"])
 apply (rotate_tac -2,
        frule_tac a = "Suc n" in forall_spec, simp) apply (
        thin_tac "∀x≤Suc n.
            ∀y≤Suc n.
               cmp (transpos (f (Suc n)) (Suc n)) f x =
               cmp (transpos (f (Suc n)) (Suc n)) f y -->
               x = y") apply (
        rotate_tac -1,
        frule_tac a = l in forall_spec1) apply (
        thin_tac "∀y≤Suc n.
            cmp (transpos (f (Suc n)) (Suc n)) f (Suc n) =
            cmp (transpos (f (Suc n)) (Suc n)) f y -->
            Suc n = y")
apply (metis Nfunc_iNJTr comp_transpos_1 le_SucE le_SucI le_refl less_Suc_eq_le transpos_ij_2)
apply (simp add:isom_gch_unitsTr1_4)
apply (simp add:isom_Gchains_def[of "n"])
apply (rule allI, rule impI)
apply (simp add:cmp_def)
apply (cut_tac l = "f i" in transpos_mem[of "f (Suc n)" "Suc n" "Suc n"])
       apply simp+
apply (cut_tac k = "f i" in cmp_transpos1 [of "f (Suc n)" "Suc n" "Suc n"])
       apply simp+
 apply (simp add:cmp_def)
apply (thin_tac "transpos (f (Suc n)) (Suc n) (transpos (f (Suc n)) (Suc n) (f i)) = f i")
apply (simp add:isom_Gchains_def)
done

lemma isom_gch_unitsTr1_6:"[|Ugp E; f (Suc n) ≠ Suc n; Gchain (Suc n) g ∧ 
       Gchain (Suc n) h ∧ Gch_bridge (Suc n) g h f|]  ==> Gchain (Suc n) g ∧
          Gchain (Suc n) (cmp h (transpos (f (Suc n)) (Suc n))) ∧
          Gch_bridge (Suc n) g (cmp h (transpos (f (Suc n)) (Suc n)))
           (cmp (transpos (f (Suc n)) (Suc n)) f)"
apply (erule conjE)+
apply simp
apply (simp add:Gch_bridge_def, frule conjunct1)
apply (frule conjunct2, fold Gch_bridge_def, erule conjE)
apply (rule conjI)
apply (thin_tac "Gchain (Suc n) g")
apply (simp add:Gchain_def, rule allI, rule impI)
apply (simp add:cmp_def)
apply (cut_tac n_in_Nsetn[of "Suc n"])
apply (cut_tac l = l in transpos_mem [of "f (Suc n)" "Suc n" "Suc n"],
       simp+)
apply (simp add:Gch_bridge_def)
apply (cut_tac n_in_Nsetn[of "Suc n"])
apply (rule conjI)
apply (rule allI, rule impI)
apply (simp add:cmp_def)
apply (rule_tac l = "f l" in transpos_mem [of "f (Suc n)" "Suc n" "Suc n"],
       simp+)
apply (rule conjI)
apply (rule isom_gch_unitsTr1_3 [of "E" "f" "n"], assumption+)
apply (simp add:isom_Gchains_def, rule allI, rule impI)
apply (simp add:cmp_def)
apply (cut_tac k = "f i" in cmp_transpos1 [of "f (Suc n)" "Suc n" "Suc n"], 
       simp+) 
apply (simp add:cmp_def)
done

lemma isom_gch_unitsTr1_7_0:"[|Gchain (Suc n) h; k ≠ Suc n; k ≤ (Suc n)|]
 ==> Gchain (Suc n) (cmp h (transpos k (Suc n)))"
apply (simp add:Gchain_def)
apply (rule allI, rule impI)
apply (simp add:cmp_def) 
apply (insert n_in_Nsetn [of "Suc n"])
apply (cut_tac l = l in transpos_mem [of "k" "Suc n" "Suc n"])
apply simp+
done

lemma isom_gch_unitsTr1_7_1:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; k ≤ (Suc n)|]
 ==> {i. i ≤ (Suc n) ∧ cmp h (transpos k (Suc n)) i ≅ E} - {k , Suc n} =
                            {i. i ≤ (Suc n) ∧ h i ≅ E} - {k , Suc n}"
apply (cut_tac n_in_Nsetn[of "Suc n"])
apply auto
apply (frule_tac x = x in transpos_id_1 [of "k" "Suc n" "Suc n"], simp+)
apply (simp add:cmp_def)
apply (simp add:cmp_def)
apply (cut_tac x = x in transpos_id_1 [of "k" "Suc n" "Suc n"], simp+)
done

lemma isom_gch_unitsTr1_7_2:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; 
            k ≤ (Suc n);  h (Suc n) ≅ E|] ==> 
                          cmp h (transpos k (Suc n)) k ≅ E"
apply (simp add:cmp_def)
apply (subst transpos_ij_1 [of "k" "Suc n" "Suc n"], simp+)
done

lemma isom_gch_unitsTr1_7_3:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; 
       k ≤ (Suc n); h k ≅ E|] ==> cmp h (transpos k (Suc n)) (Suc n) ≅ E"
apply (simp add:cmp_def)
apply (subst transpos_ij_2 [of "k" "Suc n" "Suc n"], assumption+)
apply simp apply assumption+
done

lemma isom_gch_unitsTr1_7_4:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n;  
      k ≤ (Suc n); ¬ h (Suc n) ≅ E|] ==>
                ¬  cmp h (transpos k (Suc n)) k ≅ E"
apply (rule contrapos_pp, simp+)
apply (simp add:cmp_def)
apply (insert n_in_Nsetn [of "Suc n"])
apply (simp add: transpos_ij_1 [of "k" "Suc n" "Suc n"])
done

lemma isom_gch_unitsTr1_7_5:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; 
      k ≤ (Suc n); ¬ h k ≅ E|] ==>
              ¬  cmp h (transpos k (Suc n)) (Suc n) ≅ E"
apply (rule contrapos_pp, simp+)
apply (simp add:cmp_def)
apply (insert n_in_Nsetn [of "Suc n"])
apply (simp add:transpos_ij_2 [of "k" "Suc n" "Suc n"])
done

lemma isom_gch_unitsTr1_7_6:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; 
    k ≤ (Suc n); h (Suc n) ≅ E; h k ≅ E|] ==> 
    {i. i ≤ (Suc n) ∧  h i ≅ E} = 
    insert k (insert (Suc n) ({i. i ≤ (Suc n) ∧ h i ≅ E} - {k, Suc n}))" 
apply auto
done

lemma isom_gch_unitsTr1_7_7:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; 
      k ≤ (Suc n); h (Suc n) ≅ E; ¬ h k ≅ E|] ==> 
      {i. i ≤ (Suc n) ∧  h i ≅ E} = 
       insert (Suc n) ({i. i ≤ (Suc n) ∧ h i ≅ E} - {k, Suc n})" 
apply auto
done

lemma isom_gch_unitsTr1_7_8:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; 
      k ≤ (Suc n); ¬ h (Suc n) ≅ E;  h k ≅ E|] ==> 
     {i. i ≤ (Suc n) ∧  h i ≅ E} = 
       insert k ({i. i ≤ (Suc n) ∧ h i ≅ E} - {k, Suc n})"
apply auto
done

lemma isom_gch_unitsTr1_7_9:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; 
      k ≤ (Suc n); ¬ h (Suc n) ≅ E; ¬ h k ≅ E|] ==>
      {i. i ≤ (Suc n) ∧  h i ≅ E} =
      {i. i ≤ (Suc n) ∧ h i ≅ E} - {k, Suc n}" 
apply auto
done

lemma isom_gch_unitsTr1_7:"[|Ugp E; Gchain (Suc n) h; k ≠ Suc n; 
    k ≤ (Suc n)|] ==> card {i. i ≤ (Suc n) ∧ 
    cmp h (transpos k (Suc n)) i ≅ E} =  card {i. i ≤ (Suc n) ∧ h i ≅ E}"
apply (cut_tac finite_Nset[of "Suc n"])
apply (frule isom_gch_unitsTr1_7_1 [of "E" "n" "h" "k"], assumption+)
apply (cut_tac n_in_Nsetn[of "Suc n"])
apply (case_tac "h (Suc n) ≅ E")
 apply (case_tac "h k ≅ E")
 apply (subst isom_gch_unitsTr1_7_6 [of "E" "n" "h" "k"], assumption+)
 apply (frule isom_gch_unitsTr1_7_2 [of "E" "n" "h" "k"], assumption+)
 apply (frule isom_gch_unitsTr1_7_3 [of "E" "n" "h" "k"], assumption+)
apply (frule isom_gch_unitsTr1_7_0 [of "n" "h" "k" ], assumption+)
apply (subst isom_gch_unitsTr1_7_6 [of "E" "n" "cmp h (transpos k (Suc n))" "k"], assumption+)
apply (subst card_insert_disjoint) 
apply (rule finite_subset[of "insert (Suc n) 
           ({i. i ≤ (Suc n) ∧ cmp h (transpos k (Suc n)) i ≅ E} -
           {k, Suc n})" "{i. i ≤ (Suc n)}"])
 apply (rule subsetI, simp)
 apply (erule disjE)
        apply simp apply simp apply assumption
        apply simp

apply (subst card_insert_disjoint)+
 apply (rule finite_subset[of "{i. i ≤ (Suc n) ∧ cmp h (transpos k (Suc n)) i ≅ E} - {k, Suc n}" "{i. i ≤ (Suc n)}"])
 apply (rule subsetI, simp) apply assumption
 apply simp

apply (subst card_insert_disjoint)+
 apply (rule finite_subset[of "insert (Suc n) ({i. i ≤ (Suc n) ∧ h i ≅ E} - {k, Suc n})" "{i. i ≤ (Suc n)}"])
 apply (rule subsetI, simp) apply (erule disjE, simp add:n_in_Nsetn) apply simp
 apply assumption apply simp

apply (subst card_insert_disjoint)
  apply (rule finite_subset[of "{i. i ≤ (Suc n) ∧ h i ≅ E} - {k, Suc n}"
         "{i. i ≤ (Suc n)}"])
  apply (rule subsetI, simp) apply assumption apply simp
  apply simp
     
apply (subst isom_gch_unitsTr1_7_7[of "E" "n" "h" "k"], assumption+)
apply (frule isom_gch_unitsTr1_7_0[of "n" "h" "k"], assumption+)
apply (subst isom_gch_unitsTr1_7_8[of "E" "n" "cmp h (transpos k (Suc n))" 
        "k"], assumption+)
  apply (subst cmp_def)
  apply (subst transpos_ij_2[of "k" "Suc n" "Suc n"], assumption+, simp+)
 apply (simp add:cmp_def, simp add:transpos_ij_1) 

apply (subst card_insert_disjoint)
  apply (rule finite_subset[of "{i. i ≤ (Suc n) ∧ 
         cmp h (transpos k (Suc n)) i ≅ E} - {k, Suc n}" "{i. i ≤ (Suc n)}"])
  apply (rule subsetI, simp) apply assumption apply simp
apply (subst card_insert_disjoint)
  apply (rule finite_subset[of "{i. i ≤ (Suc n) ∧ h i ≅ E} - {k, Suc n}" 
         "{i. i ≤ (Suc n)}"])
  apply (rule subsetI, simp) apply assumption apply simp
  apply simp

apply (case_tac "h k ≅ E")
apply (subst isom_gch_unitsTr1_7_8 [of "E" "n" "h" "k"], assumption+)
apply (frule isom_gch_unitsTr1_7_3 [of "E" "n" "h" "k"], assumption+) 
apply (frule isom_gch_unitsTr1_7_4 [of "E" "n" "h" "k"], assumption+)
apply (frule isom_gch_unitsTr1_7_0 [of "n" "h" "k" ], assumption+)
apply (subst isom_gch_unitsTr1_7_7 [of "E" "n" "cmp h (transpos k (Suc n))" 
                                                     "k"], assumption+)
apply (subst card_insert_disjoint)
apply (rule isom_gch_units_transpTr4, assumption+) apply simp
apply (subst card_insert_disjoint)
apply (rule isom_gch_units_transpTr4, assumption+) apply simp apply simp
apply (subst isom_gch_unitsTr1_7_9 [of "E" "n" "h" "k"], assumption+)
apply (frule_tac isom_gch_unitsTr1_7_4 [of "E" "n" "h" "k"], assumption+)
apply (frule_tac isom_gch_unitsTr1_7_5 [of "E" "n" "h" "k"], assumption+)
apply (frule isom_gch_unitsTr1_7_0 [of "n" "h" "k" ], assumption+)
apply (subst isom_gch_unitsTr1_7_9 [of "E" "n" " cmp h (transpos k (Suc n))" "k"], assumption+) apply simp
done

lemma isom_gch_unitsTr1:"Ugp E ==> ∀g. ∀h. ∀f. Gchain n g ∧ 
      Gchain n h ∧ Gch_bridge n g h f -->  card {i. i ≤ n ∧ g i ≅ E} = 
           card {i. i ≤ n ∧ h i ≅ E}"
apply (induct_tac n)
 apply (rule allI)+ apply (rule impI) apply (erule conjE)+
 apply (rule card_eq)
 apply (simp add:Gch_bridge_def)
 apply (erule conjE)+
 apply (simp add:isom_Gchains_def)
apply (rule equalityI)
apply (rule subsetI) apply (simp add:CollectI) apply (erule conjE)
 apply simp
 apply (simp add:Gchain_def)
apply (rule_tac F = "g 0" and G = "h 0" in isom_gch_unitsTr4 [of _ _ "E"],
                                                  assumption+)
apply (rule subsetI) apply (simp add:CollectI) apply (erule conjE)
 apply simp
apply (simp add:Gchain_def)
apply (frule_tac F = "g 0" and G = "h 0" in isomTr1, assumption+)
apply (rule_tac F = "h 0" and G = "g 0" in isom_gch_unitsTr4 [of _ _ "E"],
                                                  assumption+)
(***** n ******)
 apply (rule allI)+  apply (rule impI)
 (** n ** case f (Suc n) = Suc n **)
 apply (case_tac "f (Suc n) = Suc n") 
 apply (subgoal_tac "card {i. i ≤ n ∧ g i ≅ E} =  card {i. i ≤ n ∧ h i ≅ E}")
 apply (thin_tac " ∀g h f.
             Gchain n g ∧ Gchain n h ∧ Gch_bridge n g h f -->
             card {i. i ≤ n ∧ g i ≅ E} =
             card {i. i ≤ n ∧ h i ≅ E}")
apply (rule isom_tgch_unitsTr0, assumption+)
 apply (frule_tac n = n and g = g and h = h and f = f in 
            isom_gch_unitsTr1_1 [of "E"], assumption+)
 apply (rotate_tac -1) 
 apply (thin_tac "Gchain (Suc n) g ∧ Gchain (Suc n) h ∧ Gch_bridge (Suc n) g h f")
 apply blast
  (**** n **** f (Suc n) ≠ Suc n ***) 
 apply (frule_tac n = n and g = g and h = h and f = f in isom_gch_unitsTr1_5 [of "E"], assumption+) 
apply (subgoal_tac "card {i. i ≤ n ∧ g i ≅ E} = card {i. i ≤ n ∧  
                       (cmp h (transpos (f (Suc n)) (Suc n))) i ≅ E}")
prefer 2 apply blast
 apply (thin_tac "∀g h f.
             Gchain n g ∧ Gchain n h ∧ Gch_bridge n g h f -->
             card {i. i ≤ n ∧ g i ≅ E} =
             card {i. i ≤ n ∧ h i ≅ E}")
 apply (thin_tac "Gchain n g ∧
          Gchain n (cmp h (transpos (f (Suc n)) (Suc n))) ∧
          Gch_bridge n g (cmp h (transpos (f (Suc n)) (Suc n)))
           (cmp (transpos (f (Suc n)) (Suc n)) f)")
apply (subgoal_tac "cmp (transpos (f (Suc n)) (Suc n)) f (Suc n) = Suc n")
apply (frule_tac n = n and g = g and h = "cmp h (transpos (f (Suc n)) (Suc n))" and f = "cmp (transpos (f (Suc n)) (Suc n)) f"  in  isom_tgch_unitsTr0 [of "E"], assumption+)
apply (rule isom_gch_unitsTr1_6, assumption+)
 apply (thin_tac "card {i. i ≤ n ∧ g i ≅ E} = card {i. i ≤ n ∧
                          cmp h (transpos (f (Suc n)) (Suc n)) i ≅ E}")
prefer 2 
 apply (erule conjE)+
 apply (simp add:Gch_bridge_def) apply (erule conjE)+
 apply (rule isom_gch_unitsTr1_2, assumption+)
apply simp
apply (erule conjE)+
apply (rule isom_gch_unitsTr1_7, assumption+)
apply (simp add:Gch_bridge_def)
done  
 
lemma isom_gch_units:"[|Ugp E; Gchain n g; Gchain n h; Gch_bridge n g h f|] ==>  
      card {i. i ≤ n ∧ g i ≅ E} = card {i. i ≤ n ∧ h i ≅ E}"
apply (simp add:isom_gch_unitsTr1)
done

lemma isom_gch_units_1:"[|Ugp E; Gchain n g; Gchain n h; ∃f. Gch_bridge n g h f|]
 ==>  card {i. i ≤ n ∧ g i ≅ E} = card {i. i ≤ n ∧ h i ≅ E}"
apply (auto del:equalityI)
apply (simp add:isom_gch_units)
done

section "19. Jordan Hoelder theorem"

subsection "rfn_tools. tools to treat refinement of a cmpser, rtos"

lemma rfn_tool1:"[| 0 < (r::nat); (k::nat) = i * r + j; j < r |] 
                                                  ==> (k div r) = i"  
proof -
 assume p1: "0 < r" and p2: "k = i * r + j" and p3: "j < r"
 from p1 and p2 have q1: "(j + i * r) div r = i + j div r" 
 apply (simp add:div_mult_self1 [of "r" "j" "i" ]) done
 from p1 and p3 have q2: "j div r = 0"
  apply (simp add:div_if) done
 from q1 and q2 have q3:"(j + i * r) div r = i"
  apply simp done
 from q3 have q4: "(i * r + j) div r = i" apply (simp add:add_commute)
  done
 from p2 and q4 show ?thesis
  apply simp 
 done
qed

lemma pos_mult_pos:"[| 0 < (r::nat); 0 < s|] ==> 0 < r * s"
by simp

lemma rfn_tool1_1:"[| 0 < (r::nat); j < r |] 
                             ==> (i * r + j) div r = i"
apply (rule rfn_tool1 [of "r" "i * r + j" "i" "j"], assumption+)  
apply simp+
done

lemma rfn_tool2:"(a::nat) < s ==> a ≤ s - Suc 0"
 apply (rule less_le_diff) apply simp+
done

lemma rfn_tool3:"(0::nat) ≤ m ==> (m + n) - n = m"
apply auto
done

lemma rfn_tool11:"[|0 < b; (a::nat) ≤ b - Suc 0|] ==> a < b"
 apply (frule le_less_trans [of "a" "b - Suc 0" "b"])
 apply simp+
done

lemma  rfn_tool12:"[|0 < (s::nat); (i::nat) mod s = s - 1 |] ==>
                     Suc (i div s) = (Suc i) div s "
proof -
 assume p1:"0 < s" and p2:"i mod s = s - 1"
 have q1:"i div s * s + i mod s = i"
   apply (insert mod_div_equality [of "i" "s"]) 
   apply simp done
 have q2:"Suc (i div s * s + i mod s) = i div s * s + Suc (i mod s)"
  apply (insert add_Suc_right [THEN sym, of "i div s * s" "i mod s"])
  apply assumption done
 from p1 and p2 and q2 have q3:"Suc (i div s * s + i mod s) = i div s * s + s"
  apply simp done
 from q3 have q4:"Suc (i div s * s + i mod s) = Suc (i div s) * s "
  apply simp done
 from p1 and q1 and q4 show ?thesis  
 apply auto
 done
qed

lemma  rfn_tool12_1:"[|0 < (s::nat); (l::nat) mod s < s - 1 |] ==>
                     Suc (l mod s) = (Suc l) mod s "
apply (insert mod_div_equality [of "l" "s"])
apply (insert add_Suc_right [THEN sym, of "l div s * s" "l mod s"])
apply (insert mod_mult_self1  [of "Suc (l mod s)" "l div s" "s"])
apply (frule Suc_mono [of "l mod s" "s - 1"]) apply simp
done

lemma  rfn_tool12_2:"[|0 < (s::nat); (i::nat) mod s = s - Suc 0|] ==>
                                               (Suc i) mod s = 0" 
apply (insert mod_div_equality [THEN sym, of "i" "s"])
apply (insert add_Suc_right [THEN sym, of "i div s * s" "i mod s"])
apply auto
done

lemma rfn_tool13:"[| (0::nat) < r; a = b |] ==> a mod r = b mod r"
apply simp
done

lemma rfn_tool13_1:"[| (0::nat) < r; a = b |] ==> a div r = b div r"
apply simp
done

lemma div_Tr1:"[| (0::nat) < r; 0 < s; l ≤ s * r|] ==> l div s ≤ r"
 apply (frule_tac m = l and n = "s * r" and k = s in div_le_mono)
 apply simp
done

lemma div_Tr2:"[|(0::nat) < r; 0 < s; l < s * r|] ==> l div s ≤ r - Suc 0"
apply (rule contrapos_pp, simp+)
apply (simp add: not_less [symmetric, of "l div s" "r - Suc 0"])
apply (frule Suc_leI [of "r - Suc 0" "l div s"])
apply simp
apply (frule less_imp_le [of "l" "s * r"])
apply (frule div_le_mono [of "l" "s * r" "s"]) apply simp
apply (insert mod_div_equality [THEN sym, of "l" "s"])
apply (frule sym) apply (thin_tac "r = l div s")
apply simp apply (simp add:mult_commute [of "r" "s"])
done

lemma div_Tr3:"[|(0::nat) < r; 0 < s; l < s * r|] ==> Suc (l div s) ≤ r"
apply (frule_tac div_Tr2[of "r" "s"], assumption+,
       cut_tac n1 = "l div s" and m1 = "r - Suc 0" in  Suc_le_mono[THEN sym])
apply simp
done

lemma div_Tr3_1:"[|(0::nat) < r; 0 < s; l mod s = s - 1|] ==>  Suc l div s = Suc (l div s)"
apply (frule rfn_tool12 [of "s" "l"], assumption+)
 apply (rotate_tac -1) apply simp
done

lemma div_Tr3_2:"[|(0::nat) < r; 0 < s; l mod s < s - 1|] ==> 
                                       l div s = Suc l div s"
apply (frule Suc_mono [of "l mod s" "s - 1"]) apply simp
apply (cut_tac mod_div_equality [of "l" "s"])
apply (cut_tac add_Suc_right [THEN sym, of "l div s * s" "l mod s"])
apply (cut_tac zero_less_Suc[of "l mod s"],
       frule less_trans[of "0" "Suc (l mod s)" "s"], assumption+)
apply (frule rfn_tool13_1 [of "s" "Suc (l div s * s + l mod s)" "l div s * s + Suc (l mod s)"], assumption+)
apply (frule div_mult_self1 [of "s" "Suc (l mod s)" "l div s"])
apply (frule div_if [of "s" "Suc (l mod s)"], simp)
done

lemma mod_div_injTr:"[|(0::nat) < r; x mod r = y mod r; x div r = y div r|]
                      ==> x = y"
apply (cut_tac mod_div_equality[of "x" "r"])
apply simp
done

constdefs 
 rtos::"[nat, nat] => (nat => nat)"
  "rtos r s i == if i < r * s then (i mod s) * r + i div s else r * s"
 
lemma rtos_hom0:"[|(0::nat) < r; (0::nat) < s; i ≤ (r * s - Suc 0)|] ==>
   i div s < r" 
apply (frule div_Tr2 [of "r" "s" "i"], assumption+)
apply (simp add:mult_commute [of "r" "s"])
apply (rule le_less_trans, assumption+) apply simp 
apply (rule le_less_trans, assumption+) apply simp
done

lemma rtos_hom1:"[|(0::nat) < r; 0 < s; l ≤ (r * s - Suc 0)|] ==> 
 (rtos r s) l ≤ (s * r - Suc 0)"
apply (simp add:rtos_def)
apply (frule le_less_trans [of "l" "r * s - Suc 0" "r * s"])
 apply simp
 apply simp 
apply (frule mod_less_divisor [of "s" "l"])
 apply (frule less_le_diff [of "l mod s" "s"])
 apply (frule_tac i = "l mod s" and j = "s - Suc 0" and k = r and l = r in
         mult_le_mono, simp)
 apply (frule_tac i = "l mod s * r" and j = "(s - Suc 0) * r" and k = "l div s" and l = "r - Suc 0" in add_le_mono)
 apply (rule div_Tr2, assumption+, simp add:mult_commute)
 apply (simp add:diff_mult_distrib)
done

lemma rtos_hom2:"[|(0::nat) < r; (0::nat) < s; l ≤ (r * s - Suc 0)|] ==> 
 rtos r s l ≤ (r * s - Suc 0)"
apply (insert  rtos_hom1 [of "r" "s"]) apply simp
apply (simp add:mult_commute)
done

lemma rtos_hom3:"[|(0::nat) < r; 0 < s; i ≤ (r * s - Suc 0) |] ==>
                   (rtos r s i div r) = i mod s"
apply (simp add:rtos_def)
 apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"])
 apply simp apply simp
apply (subst add_commute)
 apply (frule div_mult_self1 [of "r" "i div s" "i mod s"])
 apply (frule div_Tr2 [of "r" "s" "i"], assumption+) 
 apply (simp add:mult_commute)
apply (frule le_less_trans [of "i div s" "r - Suc 0" "r"])
 apply simp
 apply (simp add:div_if [of "r" "i div s"])
done

lemma rtos_hom3_1:"[|(0::nat) < r; (0::nat) < s; i ≤ (r * s - Suc 0) |] ==>
  (rtos r s i mod  r) = i div s"
apply (simp add:rtos_def)
apply (simp add:rfn_tool11 [of "r * s" "i"])
 apply (frule rtos_hom0 [of "r" "s" "i"], assumption+)
 apply (simp add:mem_of_Nset)
 apply (subst add_commute)
 apply (subst mod_mult_self1 [of "i div s" "i mod s" "r"])
apply (subst mod_less, assumption+) apply simp 
done

lemma rtos_hom5:"[|(0::nat) < r; (0::nat) < s; i ≤ (r *s - Suc 0); 
i div s = r - Suc 0 |] ==> Suc (rtos r s i) div r = Suc (i mod s)"
 apply (frule mult_less_mono2[of "0" "s" "r"],
        simp only:nat_mult_commute,
        simp only:mult_0_right)
 apply (frule rfn_tool11 [of "r * s" "i"], assumption+)
 apply (simp add:rtos_def, simp add:div_add_self2)
done

lemma rtos_hom7:"[|(0::nat) < r; (0::nat) < s; i ≤ (r * s - Suc 0); 
                   i div s = r - Suc 0 |] ==> Suc (rtos r s i) mod r = 0"
apply (frule rtos_hom0 [of "r" "s" "i"], assumption+)
apply (simp add: rtos_def) 
apply (frule mult_less_mono2[of "0" "s" "r"],
        simp only:nat_mult_commute,
        simp only:mult_0_right)
apply (simp add:rfn_tool11 [of "r * s" "i"])
done

lemma rtos_inj:"[| (0::nat) < r; (0::nat) < s |] ==> 
                   inj_on (rtos r s) {i. i ≤ (r * s - Suc 0)}"
apply (simp add:inj_on_def) 
apply ((rule allI, rule impI)+, rule impI)
 apply (frule_tac i1 = x in rtos_hom3_1[THEN sym, of "r" "s"], assumption+,
        frule_tac i1 = x in rtos_hom3[THEN sym, of "r" "s"], assumption+,
        frule_tac i = y in rtos_hom3_1[of "r" "s"], assumption+,
        frule_tac i = y in rtos_hom3[of "r" "s"], assumption+)
 apply simp
 apply (rule_tac x = x and y = y in mod_div_injTr[of "s"], assumption+)        
done

lemma rtos_rs_Tr1:"[|(0::nat) < r; 0 < s |] ==> rtos r s (r * s) = r * s"
apply (simp add:rtos_def)
done

lemma rtos_rs_Tr2:"[|(0::nat) < r; 0 < s |] ==>
                                 ∀l ≤ (r * s). rtos r s l ≤ (r * s)"
apply (rule allI, rule impI)
 apply (case_tac "l = r * s", simp add:rtos_rs_Tr1)
 apply (frule le_imp_less_or_eq) 
 apply (thin_tac "l ≤ r * s", simp)
 apply (frule mult_less_mono2[of "0" "s" "r"],
        simp only:nat_mult_commute,
        simp only:mult_0_right)
 apply (frule_tac  r = r and s = s and l = l in rtos_hom2, assumption+)
 apply (rule less_le_diff)
 apply simp
 apply (metis le_pre_le) 
done
    
lemma rtos_rs_Tr3:"[|(0::nat) < r; 0 < s |] ==>
                             inj_on (rtos r s) {i. i ≤ (r * s)}"
apply (frule rtos_inj [of "r" "s"], assumption+)
apply (simp add:inj_on_def [of _ "{i. i ≤ (r * s)}"])
 apply ((rule allI, rule impI)+, rule impI)
 apply (case_tac "x = r * s")
 apply (rule contrapos_pp, simp+)
 apply (frule not_sym)
 apply (frule mult_less_mono2[of "0" "s" "r"],
        simp only:nat_mult_commute, simp only:mult_0_right)
 apply (cut_tac x = y and n = "r * s - Suc 0" in Nset_pre, simp+) 
 apply (frule_tac l = y in rtos_hom1[of "r" "s"], assumption+)
 apply (simp only: rtos_rs_Tr1) 
 apply (frule sym, thin_tac "r * s = rtos r s y", simp)
 apply (simp add:mult_commute[of "s" "r"])
 apply (simp add: not_less [symmetric, of "r * s" "r * s - Suc 0"])

apply (frule mult_less_mono2[of "0" "s" "r"],
       simp only:nat_mult_commute, simp only:mult_0_right,
       cut_tac x = x in Nset_pre[of _ "r * s - Suc 0"], simp+) 
 apply (case_tac "y = r * s")
        apply (simp add: rtos_rs_Tr1)
 apply (frule_tac l = x in rtos_hom1[of "r" "s"], assumption+)       
       apply (simp add:mult_commute[of "s" "r"])
        apply (simp add: not_less [symmetric, of "r * s" "r * s - Suc 0"])
 
 apply (cut_tac x = y in Nset_pre[of _ "r * s - Suc 0"], simp+)
 apply (frule rtos_inj[of "r" "s"], assumption+)
  apply (simp add:inj_on_def)
done

lemma Qw_cmpser:"[|Group G; w_cmpser G (Suc n) f |] ==> 
                               Gchain n (Qw_cmpser G f)" 
apply (simp add:Gchain_def)
apply (rule allI, rule impI)
 apply (simp  add:Qw_cmpser_def)
 apply (simp add:w_cmpser_def)
 apply (erule conjE)
 apply (simp add:d_gchain_def)
 apply (cut_tac c = l in subsetD[of "{i. i ≤ n}" "{i. i ≤ (Suc n)}"], 
        rule subsetI, simp, simp)
 apply (frule_tac H = "f l" in Group.Group_Gp[of "G"],
        frule_tac a = l in forall_spec, simp+)
 apply (frule_tac G = "Gp G (f l)" and N = "f (Suc l)" in Group.Group_Qg)
 apply blast apply assumption
done

constdefs (structure G)
 wcsr_rfns::"[_ , nat, nat => 'a set, nat] => (nat => 'a set) set"
  "wcsr_rfns G r f s  == {h. tw_cmpser G (s * r) h ∧ 
                                 (∀i ≤ r. h (i * s) = f i)}"
      (** where f ∈ tw_cmpser G r **)
  (**  0-+-+-2-+-+-4-+-+-6  h 
      f 0   f 1   f 2   f 3  **)

 trivial_rfn::"[_ , nat, nat => 'a set, nat] => (nat => 'a set)"
   "trivial_rfn G r f s k == if k < (s * r) then f (k div s) else f r"

lemma (in Group) rfn_tool8:"[|compseries G r f; 0 < r|] ==> d_gchain G r f" 
apply (simp add:compseries_def tW_cmpser_def) 
apply (erule conjE)+
apply (simp add:tD_gchain_def) apply (erule conjE)+
apply (simp add:D_gchain_is_d_gchain)
done

lemma (in Group) rfn_tool16:"[|0 < r; 0 < s; i ≤ (s * r - Suc 0); 
 G » f (i div s); (Gp G (f (i div s))) \<triangleright> f (Suc (i div s)); 
 (Gp G (f (i div s))) » (f (i div s) ∩ g (s - Suc 0))|]  ==> 
 (Gp G ((f (Suc (i div s)) \<struct>G (f (i div s) ∩ g (s - Suc 0))))) \<triangleright> 
                                                   (f (Suc (i div s)))"
apply (rule ZassenhausTr4_1 [of "f (i div s)" "f (Suc (i div s))" "g (s - Suc 0)"], assumption+)
done

text{* Show existence of the trivial refinement. This is not necessary
to prove JHS *}

lemma rfn_tool30:"[|0 < r; 0 < s; l div s * s + s < s * r|] 
                ==> Suc (l div s) < r"
apply (simp add:mult_commute[of "s" "r"])
apply (cut_tac add_mult_distrib[THEN sym, of "l div s" "s" "1"])
      apply (simp only:nat_mult_1)
apply (thin_tac "l div s * s + s = (l div s + 1) * s")
apply (cut_tac mult_less_cancel2[of "l div s + 1" "s" "r"])
       apply simp
done

lemma (in Group) simple_grouptr0:"[|G » H; G \<triangleright> K; K ⊆ H; simple_Group (G / K)|]
  ==> H = carrier G ∨ H = K" 
apply (simp add:simple_Group_def)
apply (frule subg_Qsubg[of "H" "K"], assumption+) 
apply (rotate_tac -1)
apply (frule in_set_with_P[of _ "carrier ((Gp G H) / K)"])
       apply simp
       apply (thin_tac "{N. G / K » N } = {carrier (G / K), {\<one>G / K}}")
 apply (erule disjE) 
 apply (frule sg_subset[of "H"])
 apply (frule equalityI[of "H" "carrier G"])
 apply (rule subsetI)
 apply (simp add:Qg_carrier)
 apply (frule nsg_sg[of "K"])
 apply (frule_tac a = x in rcs_in_set_rcs[of "K"], assumption+)
 apply (frule sym, thin_tac "carrier ((\<natural>H)/ K) = set_rcs G K", simp,
        thin_tac "set_rcs G K = carrier ((\<natural>H) / K)")
 apply (frule Group_Gp[of "H"], simp add:Group.Qg_carrier[of "Gp G H" "K"],
        simp add:set_rcs_def, erule bexE, simp add:Gp_carrier)

 apply (simp add:rcs_in_Gp[THEN sym])
 apply (frule_tac a = x in a_in_rcs[of "K"], assumption+, simp,
        thin_tac "K • x = K • a",
        thin_tac "G / K » {C. ∃a∈H. C = K • a}",
        simp add:rcs_def, erule bexE,
        frule_tac c = h in subsetD[of "K" "H"], assumption+)
 apply (frule_tac x = h and y = a in sg_mult_closed[of "H"], assumption+, simp)
 apply simp
 
apply (frule equalityI[THEN sym, of "K" "H"]) 
 apply (rule subsetI)
 apply (frule Group_Gp[of "H"], simp add:Group.Qg_carrier)
 apply (simp add:Qg_one)
 apply (frule nsg_sg[of "K"])
 apply (frule_tac a = x in Group.rcs_in_set_rcs[of "Gp G H" "K"]) 
  apply (simp add:sg_sg) apply (simp add:Gp_carrier)
  apply simp
  apply (frule_tac a = x in Group.rcs_fixed[of "Gp G H" "K"])
   apply (simp add:sg_sg) apply (simp add:Gp_carrier) apply assumption+
 apply simp
done

lemma (in Group) compser_nsg:"[|0 < n; compseries G n f; i ≤ (n - 1)|]
                  ==>  Gp G (f i) \<triangleright> (f (Suc i))"
apply (simp add:compseries_def tW_cmpser_def)
done

lemma (in Group) compseriesTr5:"[|0 < n; compseries G n f; i ≤ (n - Suc 0)|]
          ==> (f (Suc i)) ⊆  (f i)"
apply (frule compseriesTr4[of "n" "f"])
apply (frule w_cmpser_is_d_gchain[of "n" "f"])
apply (simp add:d_gchain_def,
     cut_tac n_in_Nsetn[of "n"],
     frule_tac a = n in forall_spec, simp,
     thin_tac "∀l ≤ n. G » f l  ∧ (∀l ≤ (n - Suc 0). f (Suc l) ⊆ f l)",
     erule conjE, blast)
done

lemma (in Group) refine_cmpserTr0:"[|0 < n; compseries G n f; i ≤ (n - 1);
        G » H;  f (Suc i) ⊆ H ∧ H ⊆ f i|] ==> H = f (Suc i) ∨ H = f i"
apply (frule compseriesTr0 [of "n" "f" "i"]) 
 apply (cut_tac lessI[of "n - Suc 0"], simp only:Suc_pred, simp)
 apply (frule Group_Gp [of "f i"]) 
 apply (erule conjE)
 apply (frule compseriesTr4[of "n" "f"])
 apply (frule w_cmpser_ns[of "n" "f" "i"], simp+) 
 apply (unfold compseries_def, frule conjunct2, fold compseries_def, simp)
 apply (frule_tac a = i in forall_spec1, 
        thin_tac "∀i≤n - Suc 0. simple_Group ((\<natural>f i) / f (Suc i))",
        simp)
 apply (frule Group.simple_grouptr0 [of "Gp G (f i)" "H" "f (Suc i)"])
    apply (simp add:sg_sg) apply assumption+ 
 apply (simp add:Gp_carrier)
apply blast
done

lemma div_Tr4:"[| (0::nat) < r; 0 < s; j < s * r |] ==> j div s * s + s ≤ r * s"
apply (frule div_Tr2 [of "r" "s" "j"], assumption+)
apply (frule mult_le_mono [of "j div s" "r - Suc 0" "s" "s"])
 apply simp
 apply (frule add_le_mono [of "j div s * s" "(r - Suc 0) * s" "s" "s"])
 apply simp
 apply (thin_tac "j div s * s ≤ (r - Suc 0) * s")
 apply (cut_tac add_mult_distrib[THEN sym, of "r - Suc 0" "s" "1"]) 
        apply (simp only:nat_mult_1)
apply simp
done

lemma (in Group) compseries_is_tW_cmpser:"[|0 < r; compseries G r f|] ==>
        tW_cmpser G r f"
by (simp add:compseries_def)

lemma (in Group) compseries_is_td_gchain:"[|0 < r; compseries G r f|] ==>
        td_gchain G r f"
apply (frule compseries_is_tW_cmpser, assumption+)
apply (simp add:tW_cmpser_def, erule conjE)
apply (thin_tac "∀l≤(r - Suc 0). (\<natural>f l) \<triangleright> (f (Suc l))")
apply (simp add:tD_gchain_def, (erule conjE)+)
apply (frule D_gchain_is_d_gchain)
apply (simp add:td_gchain_def)
done

lemma (in Group) compseries_is_D_gchain:"[|0 < r; compseries G r f|] ==>
             D_gchain G r f"
apply (frule compseriesTr1)
apply (frule tW_cmpser_is_W_cmpser)
apply (rule W_cmpser_is_D_gchain, assumption)
done

lemma divTr5:"[|0 < r; 0 < s; l < (r * s)|]  ==>
                 l div s * s ≤ l ∧ l  ≤ (Suc (l div s)) * s"
apply (insert mod_div_equality [THEN sym, of "l" "s"])
apply (rule conjI)
apply (insert le_add1 [of "l div s * s" "l mod s"])
apply simp
apply (frule mod_less_divisor [of "s" "l"])
 apply (frule less_imp_le [of "l mod s" "s"]) 
 apply (insert self_le [of "l div s * s"])
 apply (frule add_le_mono [of "l div s * s" "l div s * s" "l mod s" "s"])
 apply assumption+
apply (thin_tac "l div s * s ≤ l div s * s + l mod s")
apply (thin_tac "l div s * s ≤ l div s * s")
apply simp
done  

lemma (in Group) rfn_compseries_iMTr1:"[|0 < r; 0 < s; compseries G r f; 
h ∈ wcsr_rfns G r f s|] ==>  f ` {i. i ≤ r} ⊆  h ` {i. i ≤ (s * r)}"
apply (simp add:wcsr_rfns_def) apply (rule subsetI)
apply (simp add:image_def)
apply (auto del:equalityI)
apply (frule_tac i = xa in mult_le_mono [of _ "r" "s" "s"])
apply simp
apply (simp add:mult_commute [of "r" "s"])
apply (frule_tac a = xa in forall_spec, assumption,
       thin_tac "∀i≤r. h (i * s) = f i")
apply (frule sym, thin_tac "h (xa * s) = f xa")
 apply (cut_tac a = xa in mult_mono[of _ r s s], simp, simp, simp, simp) 
 apply (simp only:mult_commute[of r s], blast) 
done

lemma rfn_compseries_iMTr2:"[|0 < r; 0 < s; xa < s * r |] ==>
         xa div s * s ≤ r * s ∧ Suc (xa div s) * s ≤ r * s"
apply (frule div_Tr1 [of "r" "s" "xa"], assumption+)
 apply (simp add:less_imp_le)
apply (rule conjI)
 apply (simp add:mult_le_mono [of "xa div s" "r" "s" "s"])
apply (thin_tac "xa div s ≤ r")
apply (frule div_Tr2[of "r" "s" "xa"], assumption+)
 apply (thin_tac "xa < s * r")
 apply (frule le_less_trans [of "xa div s" "r - Suc 0" "r"])
 apply simp 
 apply (frule Suc_leI [of "xa div s" "r"])
 apply (rule mult_le_mono [of "Suc (xa div s)" "r" "s" "s"], assumption+)
apply simp
done

lemma (in Group) rfn_compseries_iMTr3:"[|0 < r; 0 < s; compseries G r f; 
      j ≤ r; ∀i ≤ r. h (i * s) = f i|] ==>  h (j * s) = f j"
apply blast
done

lemma (in Group) rfn_compseries_iM:"[|0 < r; 0 < s; compseries G r f; 
      h ∈ wcsr_rfns G r f s|]  ==> card (h `{i. i ≤ (s * r)}) = r + 1"
apply (frule compseries_is_D_gchain, assumption+)
apply (frule D_gchain1)
 apply (simp add:finite_Nset)
 apply (subst card_Nset[THEN sym, of "r"])
 apply (subst card_image[THEN sym, of "f" "{i. i ≤ r}"], assumption+)
 apply (rule card_eq[of "h ` {i. i ≤ (s * r)}" "f ` {i. i ≤ r}"])
apply (frule rfn_compseries_iMTr1[of "r" "s" "f" "h"], assumption+)
apply (rule equalityI[of "h ` {i. i ≤ (s * r)}" "f ` {i. i ≤ r}"])
prefer 2 apply simp
apply (rule subsetI,
       thin_tac "f ` {i. i ≤ r} ⊆ h ` {i. i ≤ s * r}")
apply (frule_tac b = x and f = h and A = "{i. i ≤ (s * r)}" in mem_in_image3,
       erule bexE) apply (simp add:mult_commute[of "s" "r"])
apply (simp add:wcsr_rfns_def, (erule conjE)+)
apply (frule rfn_compseries_iMTr3[of "r" "s" "f" "r" "h"], assumption+,
       simp add:n_in_Nsetn, assumption+, subst image_def, simp)
 apply (case_tac "a = s * r", simp add:mult_commute[of "s" "r"],
        cut_tac n_in_Nsetn[of "r"], blast)
 apply (simp add:mult_commute[of "s" "r"])
 apply (frule_tac m = a and n = "r * s" in noteq_le_less, assumption+)
 apply (frule tw_cmpser_is_w_cmpser, frule w_cmpser_is_d_gchain)        
 apply (frule_tac xa = a in rfn_compseries_iMTr2[of "r" "s"], assumption+)
 apply (simp add:mult_commute)
 apply (erule conjE) 
 apply (frule_tac l = a in divTr5[of "r" "s"], assumption+) 
 apply (frule pos_mult_pos[of "r" "s"], assumption+) 
 apply (erule conjE,
        frule_tac l = "a div s * s" and j = a in d_gchainTr2[of "r * s" "h"],
        assumption+)
 apply (frule_tac l = a and j = "Suc (a div s) * s" in d_gchainTr2[of "r * s" 
       "h"], assumption+)
 apply (frule_tac i = a in rtos_hom0[of "r" "s"], assumption+)
 apply (rule less_le_diff)
 apply simp
 apply (frule_tac x = "a div s" and y = r in less_imp_le,
        frule_tac a = "a div s" in forall_spec, assumption,
        frule_tac a = "Suc (a div s)" in forall_spec)
 apply (cut_tac m = "Suc (a div s)" and k = s and n = r in mult_le_cancel2)
        apply simp
 apply (thin_tac "∀i ≤ r. h (i * s) = f i") apply simp
 apply (cut_tac i = "a div s" and H = "h a"in refine_cmpserTr0[of "r" "f"],
        simp, assumption+,
        cut_tac x = "a div s" and n = r in less_le_diff, assumption, simp)
 apply (simp add:d_gchain_mem_sg[of "r * s" "h"])
 apply blast
 apply (erule disjE)
 apply (frule_tac m = "a div s" and n = r in Suc_leI, blast)
 apply (frule_tac x = "a div s" and y = r in less_imp_le, blast)
done

constdefs (structure G)
 cmp_rfn::"[_ , nat, nat => 'a set, nat, nat => 'a set] => (nat => 'a set)"
  "cmp_rfn G r f s g  == λi. (if i < s * r then  
      f (Suc (i div s)) \<struct>G (f (i div s) ∩ g (i mod s)) else {\<one>})"

 (** refinement of compseries G r f by a compseries G s g **) 

lemma (in Group) cmp_rfn0:"[|0 < r; 0 < s; compseries G r f; compseries G s g; 
 i ≤ (r - 1); j ≤ (s - 1)|] ==> G » f (Suc i) \<struct>G ((f i ) ∩ (g j))"
apply (rule ZassenhausTr2_1[of "f i" "f (Suc i)" "g j"], simp,
       rule compseriesTr0[of "r" "f" "i"], assumption+,
       frule_tac le_less_trans[of i "r - Suc 0" r], simp+)
 apply (rule compseriesTr0[of "r" "f" "Suc i"], assumption+, arith) 
 apply(rule compseriesTr0[of "s" "g" "j"], assumption+, simp);
apply (frule compseries_is_tW_cmpser[of "r" "f"], assumption+,
       simp add:tW_cmpser_def)
done

lemma (in Group) cmp_rfn1:"[|0 < r; 0 < s; compseries G r f; compseries G s g|]
  ==> f (Suc 0) \<struct>G ((f 0 ) ∩ (g 0)) = carrier G"
apply (frule compseriesTr2 [of "r" "f"])
apply (frule compseriesTr2 [of "s" "g"])
apply (frule compseriesTr0 [of _ "f" "Suc 0"])
apply simp 
apply simp
apply (rule K_absorb_HK[of "f (Suc 0)" "carrier G"], assumption+,
       simp add:special_sg_G)
apply (rule sg_subset, assumption)
done

lemma (in Group) cmp_rfn2:"[|0 < r; 0 < s; compseries G r f; compseries G s g;
           l ≤ (s * r)|]  ==> G » cmp_rfn G r f s g l"
 apply (simp add:cmp_rfn_def)
 apply (case_tac "l < s * r")
 apply simp
 apply (frule_tac i = "l div s" and j = "l mod s" in cmp_rfn0 [of "r" "s"],
                                                 assumption+)
 apply (simp add:div_Tr2)
 apply (frule_tac m = l in mod_less_divisor [of "s"])
 apply (frule_tac x = "l mod s" and n = s in less_le_diff)
 apply simp apply assumption 
apply simp
 apply (rule special_sg_e)
done

lemma (in Group) cmp_rfn3:"[|0 < r; 0 < s; compseries G r f; compseries G s g|]
 ==> cmp_rfn G r f s g 0 = carrier G ∧ cmp_rfn G r f s g (s * r) = {\<one>}"
apply (rule conjI) 
 apply (simp add:cmp_rfn_def)
 apply (rule cmp_rfn1 [of "r" "s" "f" "g"], assumption+)
 apply (simp add:cmp_rfn_def)
done
 
lemma rfn_tool20:"[|(0::nat) < m; a = b * m + c; c < m |] ==>  a mod m = c"
apply (simp add:add_commute)
done

lemma Suci_mod_s_2:"[|0 < r; 0 < s; i ≤ r * s - Suc 0; i mod s < s - Suc 0|]
        ==> (Suc i) mod s = Suc (i mod s)"
apply (insert mod_div_equality [of "i" "s", THEN sym])
apply (subgoal_tac "Suc i = i div s * s + Suc (i mod s)")
apply (subgoal_tac "Suc i mod s  = (i div s * s + Suc (i mod s)) mod s")
 apply (subgoal_tac "Suc (i mod s) < s")
 apply (frule_tac m = s and a = "Suc i" and b = "i div s" and c = "Suc (i mod s)" in rfn_tool20, assumption+)
 apply (subgoal_tac "Suc (i mod s) < Suc (s - Suc 0)")  apply simp
 apply (simp del:Suc_pred)
 apply simp+
done

lemma (in Group) inter_sgsTr1:"[|0 < r; 0 < s; compseries G r f; compseries G s g; i < r * s|]  ==> G » f (i div s) ∩ g (s - Suc 0)"
 apply (rule inter_sgs)
 apply (rule compseriesTr0, assumption+)  
 apply (frule less_imp_le [of "i" "r * s"])
 apply (frule div_Tr1 [of "r" "s" "i"], assumption+) 
 apply (simp add:mult_commute, simp)
 apply (rule compseriesTr0, simp+) 
done

lemma (in Group) JHS_Tr0_2:"[|0 < r; 0 < s; compseries G r f; compseries G s g|]
==> ∀i ≤ (s * r - Suc 0). Gp G (cmp_rfn G r f s g i) \<triangleright> 
                                           cmp_rfn G r f s g (Suc i)"

apply (frule compseriesTr4 [of "r" "f"], frule compseriesTr4 [of "s" "g"])
apply (rule allI, rule impI)
 apply (frule pos_mult_pos [of "s" "r"], assumption+,
        frule_tac a = i in  rfn_tool11 [of "s * r"], assumption+,
        frule_tac l = i in div_Tr1 [of "r" "s"], assumption+,
        simp add:less_imp_le)
 apply (frule_tac x = "i div s" in mem_of_Nset [of _ "r"],
        thin_tac "i div s ≤ r", thin_tac "i ≤ s * r - Suc 0",
        frule_tac l = i in div_Tr2 [of "r" "s"], assumption+,
        frule_tac x = "i div s" in mem_of_Nset [of _ "r - Suc 0"],
        frule_tac a = "i div s" in rfn_tool11 [of "r"], assumption+,
        frule_tac m = "i div s" in Suc_leI [of _ "r"],
        frule_tac x = "Suc (i div s)" in mem_of_Nset [of _ "r"],
        thin_tac "i div s < r", thin_tac "Suc (i div s) ≤ r")
 apply (simp add:cmp_rfn_def)
 apply (case_tac "Suc i < s * r", simp,
         case_tac "i mod s = s - 1",
         cut_tac l = i in div_Tr3_1 [of "r" "s"],
         simp+,  
         cut_tac l = "Suc i" in div_Tr2 [of "r" "s"], simp+,
         cut_tac x = "Suc i div s" in mem_of_Nset [of _ "r - Suc 0"],
         simp,
         cut_tac a = "Suc i div s" in  rfn_tool11 [of "r"], simp+,
         cut_tac x = "Suc i div s" in less_mem_of_Nset [of _ "r"], simp,
         cut_tac m = "Suc i div s" in Suc_leI [of _ "r"], simp,
         frule_tac x = "Suc (Suc i div s)" in mem_of_Nset [of _ "r"],
         frule w_cmpser_is_d_gchain [of"r" "f"],
         frule_tac rfn_tool12_2 [of "s"], assumption+,
         thin_tac "i mod s = s - Suc 0",
         thin_tac "Suc i div s ∈ {i. i ≤ r}",
         thin_tac "Suc (Suc i div s) ≤ r")
  apply (simp,
         cut_tac l = "i div s" and j = "Suc (i div s)" in 
                       d_gchainTr2 [of "r" "f"], simp, assumption+,
         cut_tac x = "i div s" and y = "Suc (i div s)" and z = r in 
          less_trans, simp, assumption, simp add:less_imp_le, 
          simp add:less_imp_le, simp,
         cut_tac l = "Suc (i div s)" and j = "Suc (Suc (i div s))" in 
                       d_gchainTr2 [of "r" "f"], simp+,
         thin_tac "Suc i div s = Suc (i div s)", thin_tac "Suc i mod s = 0",
         cut_tac i = "i div s" in compseriesTr0 [of "r" "f"], assumption,
         simp, 
         cut_tac i = "Suc (i div s)" in compseriesTr0 [of  "r" "f"],
               assumption, simp add:less_imp_le,
         cut_tac i = "Suc (Suc (i div s))" in compseriesTr0 [of "r" "f"],
               assumption, simp)  
                                                            
 apply (subst compseriesTr2 [of "s" "g"], assumption, 
        frule_tac H = "f (i div s)" in sg_subset,
        frule_tac H = "f (Suc (i div s))" in sg_subset,
        frule_tac A = "f (Suc (i div s))" in Int_absorb2 [of _ "carrier G"], 
        simp,
        thin_tac "f (Suc (i div s)) ∩ carrier G = f (Suc (i div s))",
        frule_tac H = "f (Suc (Suc (i div s)))" and K = "f (Suc (i div s))" in
        K_absorb_HK, assumption+,
        simp, 
        thin_tac "f (Suc (Suc (i div s))) \<struct>G (f (Suc (i div s))) = 
                                                         f (Suc (i div s))")
apply (rule rfn_tool16 [of "r" "s" _], simp+,
       cut_tac x = "Suc i" and y = "s * r" and z = "Suc (s *r)" in 
       less_trans, assumption, simp,
       thin_tac "Suc i < s * r", simp);
apply (rule compser_nsg[of r f], simp+) 
 apply (rule_tac  H = "f (i div s) ∩ g (s - Suc 0)" and K = "f (i div s)" in 
        sg_sg, assumption+,
        cut_tac i = i in inter_sgsTr1 [of r s f g], simp+,
        cut_tac x = i and y = "Suc i" and z = "s * r" in less_trans,
        simp+,
        cut_tac x = i and y = "Suc i" and z = "r * s" in less_trans,
        simp+,
        simp add:mult_commute, assumption+,
        simp add:Int_lower1)
 apply (frule_tac m = i in mod_less_divisor [of "s"],
        frule_tac x = "i mod s" in less_le_diff [of _ "s"],
        simp, 
        frule_tac m = "i mod s" in noteq_le_less [of _ "s - Suc 0"], 
        assumption+,
        thin_tac "i mod s ≠ s - Suc 0", thin_tac "i mod s ≤ s - Suc 0");
     apply (frule_tac x = "i mod s" and y = "s - Suc 0" and z = s in 
            less_trans, simp,
        frule_tac k = "i mod s" in nat_pos2 [of _ s],
        cut_tac l1 = i in div_Tr3_2 [THEN sym, of "r" "s"], simp+,
        frule_tac i = "i mod s" in compser_nsg [of "s" "g"], assumption+,
        simp, 
        frule_tac i = "Suc (i div s)" in compseriesTr0 [of "r" "f"], 
                  assumption+,
        frule_tac m = "i mod s" in Suc_mono [of _ "s - Suc 0"],
        simp only:Suc_pred)
 apply (frule_tac x = "Suc (i mod s)" in less_mem_of_Nset [of _ "s"],
        cut_tac i = "Suc (i mod s)" in compseriesTr0 [of "s" "g"], simp+,
        cut_tac H = "f (i div s)" and ?H1.0 = "f (Suc (i div s))" and 
        K = "g (i mod s)" and ?K1.0 = "g (Suc (i mod s))" in ZassenhausTr3,
        rule_tac i = "i div s" in compseriesTr0 [of "r" "f"], assumption+,
        frule_tac x = "i div s" and y = "r - Suc 0" and z = r in
        le_less_trans, simp, simp, assumption+,
        frule_tac i = "i mod s" in compseriesTr0 [of "s" "g"], 
        simp, simp, simp)
apply (rule_tac i = "i div s" in compser_nsg[of r f], simp+,
       cut_tac i = i in Suci_mod_s_2[of r s], simp+,
       cut_tac x = "Suc i" and y = "s * r" and z = "Suc (s *r)" in 
       less_trans, assumption, simp,
       thin_tac "Suc i < s * r", simp,
       frule_tac x = i and n = "s * r" in less_le_diff,
       simp add:mult_commute[of s r], simp+)
apply (cut_tac a = "i mod s" in rfn_tool11 [of "s"], simp+,
       frule_tac m = i in mod_less_divisor [of "s"],
       frule_tac x = "i mod s" in less_le_diff [of _ "s"], simp,
       rule special_nsg_e,
       cut_tac i = "i div s" and j = "i mod s" in cmp_rfn0[of r s f g],
       simp+)
done

lemma (in Group) cmp_rfn4:"[|0 < r; 0 < s; compseries G r f; 
      compseries G s g; l ≤ (s * r - Suc 0)|] ==>
                 cmp_rfn G r f s g (Suc l) ⊆ cmp_rfn G r f s g l"
apply (frule JHS_Tr0_2 [of "r" "s" "f" "g"], assumption+)
apply (frule_tac a = l in forall_spec, simp,
       thin_tac "∀i ≤ (s * r - Suc 0).
        (\<natural>(cmp_rfn G r f s g i)) \<triangleright> (cmp_rfn G r f s g (Suc i))")
apply (frule cmp_rfn2 [of "r" "s" "f" "g" "l"], assumption+,
       frule le_less_trans [of "l" "s * r - Suc 0" "s* r"], simp,
       simp add:less_imp_le)
apply (frule Group_Gp [of "cmp_rfn G r f s g l"],
       frule Group.nsg_subset [of "Gp G (cmp_rfn G r f s g l)"  ], 
       assumption+, simp add:Gp_carrier)
done

lemma (in Group) cmp_rfn5:"[|0 < r; 0 < s; compseries G r f; compseries G s g|]
    ==> ∀i ≤ r. cmp_rfn G r f s g (i * s) = f i"
apply (rule allI, rule impI)
apply (simp add:cmp_rfn_def)
apply (case_tac "i < r", simp,
       frule_tac x = i in less_imp_le [of _ "r"],
       frule_tac x = i in mem_of_Nset[of _ "r"],
       frule_tac i = i in compseriesTr0 [of "r" "f"], assumption+,
       thin_tac "i ≤ r",
       frule_tac m = i in Suc_leI [of _ "r"],
       frule_tac x = "Suc i" in mem_of_Nset[of _ "r"],
       frule_tac i = "Suc i" in compseriesTr0 [of "r" "f"], assumption+)
apply (subst compseriesTr2 [of "s" "g"], assumption+,
       frule_tac H = "f i" in sg_subset,
       subst Int_absorb2, assumption+,
       frule rfn_tool8, assumption+, simp)
apply (cut_tac n = i in zero_less_Suc,
       frule_tac x = 0 and y = "Suc i" and z = r in less_le_trans, assumption+,
       frule_tac l = i and j = "Suc i" in d_gchainTr2 [of "r" "f"],
       frule compseries_is_D_gchain[of "r" "f"], assumption,
        rule D_gchain_is_d_gchain, assumption+,
       cut_tac x = i and y = "Suc i" and z = r in less_le_trans, simp+,
       rule_tac K = "f i" and H = "f (Suc i)" in  K_absorb_HK, assumption+,
       simp, frule compseries_is_td_gchain, assumption+,
       simp add:td_gchain_def) 
done

lemma (in Group) JHS_Tr0:"[|(0::nat) < r; 0 < s; compseries G r f; 
       compseries G s g|] ==> cmp_rfn G r f s g ∈ wcsr_rfns G r f s"
apply (simp add:wcsr_rfns_def,
       frule cmp_rfn5 [of "r" "s" "f" "g"], assumption+,
       simp,
       thin_tac "∀i≤r. cmp_rfn G r f s g (i * s) = f i")
apply (simp add:tw_cmpser_def, rule conjI)
prefer 2 apply (rule  JHS_Tr0_2 [of "r" "s" "f" "g"], assumption+)
apply (simp add:td_gchain_def, rule conjI)
prefer 2
apply (rule cmp_rfn3, assumption+) 
apply (simp add:d_gchain_def,
       rule allI, rule impI, rule conjI, rule cmp_rfn2, assumption+,
       rule allI, rule impI, rule cmp_rfn4, assumption+)
done

lemma rfn_tool17:"(a::nat) = b ==> a - c = b - c"
by simp

lemma isom4b:"[|Group G; G \<triangleright> N; G » H|] ==> 
                  (Gp G (N \<struct>G H) / N) ≅ (Gp G H / (H ∩ N))"
apply (frule isom4 [of "G" "N" "H"], assumption+)
apply (rule isomTr1)
 apply (simp add:Group.homom4_2)
 apply (frule Group.smult_sg_nsg[of "G" "H" "N"], assumption+)
 apply (simp add:Group.smult_commute_sg_nsg[THEN sym, of "G" "H" "N"])
apply (rule Group.homom4Tr1, assumption+)
done

lemma Suc_rtos_div_r_1:"[| 0 < r; 0 < s; i ≤ r * s - Suc 0; 
         Suc (rtos r s i) < r * s; i mod s = s - Suc 0; 
         i div s < r - Suc 0|] ==>  Suc (rtos r s i) div r = i mod s"
apply (simp add:rtos_def,
       frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp)
apply simp
apply (subgoal_tac "Suc ((s - Suc 0) * r + i div s) div r = 
                      ((s - Suc 0) * r + Suc (i div s)) div r",
       simp del:add_Suc add_Suc_right)
apply (rule rfn_tool1_1, assumption+)
apply (subgoal_tac "Suc (i div s) < Suc (r - Suc 0)")
apply (simp, simp del:Suc_pred, simp) 
done

lemma Suc_rtos_mod_r_1:"[|0 < r; 0 < s; i ≤ r * s - Suc 0; Suc (rtos r s i) < r * s; i mod s = s - Suc 0; i  div s < r - Suc 0|]
         ==> Suc (rtos r s i) mod r = Suc (i div s)"
apply (simp add:rtos_def)
 apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp)
 apply simp
 apply (subgoal_tac "Suc ((s - Suc 0) * r + i div s) mod r =
   ((s - Suc 0) * r + Suc (i div s)) mod r")
  apply (simp del:add_Suc add_Suc_right)
  apply (subgoal_tac "Suc (i div s) < r")
 apply (subst rfn_tool20 [of "r" _ "(s - Suc 0)" "Suc (i div s)"], assumption+)
 apply (simp, assumption, simp)
 apply (subgoal_tac "Suc (i div s) < Suc (r - Suc 0)", simp)
 apply (simp del:Suc_pred, simp)
done

lemma i_div_s_less:"[|0 < r; 0 < s; i ≤ r * s - Suc 0; Suc (rtos r s i) < r * s;
i mod s = s - Suc 0; Suc i < s * r |]  ==> i div s < r - Suc 0"
apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp)
apply (frule_tac  r = r and s = s and l = i in div_Tr2, assumption+)
 apply (simp add:mult_commute) 
apply (rule contrapos_pp, simp+, 
        subgoal_tac "i div s = r - Suc 0",
        thin_tac "i div s = r - Suc 0")
apply (simp add:rtos_def,
       subgoal_tac "(s - Suc 0) * r + r = r * s", simp)
 apply (thin_tac "(s - Suc 0) * r + r < r * s")
 apply (simp add:mult_commute, simp add:diff_mult_distrib2)
apply simp
done

lemma rtos_mod_r_1:"[| 0 < r; 0 < s; i ≤ r * s - Suc 0; rtos r s i < r * s;
  i mod s = s - Suc 0 |]  ==> rtos r s i mod r = i div s"
apply (simp add:rtos_def)
 apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp)
 apply simp
 apply (rule rfn_tool20, assumption+, simp)
 apply (frule_tac r = r and s = s and l = i in div_Tr2, assumption+)
 apply (rule le_less_trans, assumption+, simp add:mult_commute)
 apply (rule le_less_trans, assumption+, simp)
done

lemma Suc_i_mod_s_0_1:"[|0 < r; 0 < s; i ≤ r * s - Suc 0; i mod s = s - Suc 0|]
        ==> Suc i mod s = 0"
apply (insert mod_div_equality [of "i" "s", THEN sym])
apply simp
 apply (thin_tac "i mod s = s - Suc 0")
apply (subgoal_tac "Suc i mod s = Suc (i div s * s + s - Suc 0) mod s")
 apply (thin_tac "i = i div s * s + s - Suc 0", simp del:add_Suc)
 apply (subgoal_tac "Suc (i div s * s + s - Suc 0) mod s = (i div s * s + s) mod s") 
 apply (simp del:add_Suc)
 apply (subgoal_tac "Suc (i div s * s + s - Suc 0) = i div s * s + s")
 apply (simp del:add_Suc)
apply (rule Suc_pred [of "i div s * s + s"], simp)
done

(* same as div_Tr3_2 *)
lemma Suci_div_s_2:"[|0 < r; 0 < s; i ≤ r * s - Suc 0; i mod s < s - Suc 0|]
         ==> Suc i div s = i div s"
apply (rule div_Tr3_2 [THEN sym], assumption+)
 apply simp
done

lemma rtos_i_mod_r_2:"[|0 < r; 0 < s; i ≤ r * s - Suc 0|] ==> rtos r s i mod r = i div s"
apply (simp add:rtos_def)
apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp)
apply simp
apply (frule_tac  r = r and s = s and l = i in div_Tr2, assumption+)
 apply (simp add:mult_commute)
apply (subgoal_tac "i div s < r") 
 apply (rule rfn_tool20, assumption+, simp)
 apply assumption
 apply (rule le_less_trans, assumption+, simp)
done

lemma Suc_rtos_i_mod_r_2:"[|0 < r; 0 < s; i ≤ r * s - Suc 0; 
       i div s = r - Suc 0|] ==> Suc (rtos r s i) mod r = 0"
apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp)
apply (simp add:rtos_def)
done

lemma Suc_rtos_i_mod_r_3:"[|0 < r; 0 < s; i ≤ r * s - Suc 0; 
      i div s < r - Suc 0|] ==> Suc (rtos r s i) mod r = Suc (i div s)"
apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp)
apply (simp add:rtos_def)
apply (subst add_Suc_right[THEN sym, of "i mod s * r" "i div s"])
 apply (subst add_commute[of "i mod s * r" "Suc (i div s)"])
 apply (subst mod_mult_self1[of "Suc (i div s)" "i mod s" "r"])
 apply (cut_tac lessI[of "r - Suc 0"], simp only:Suc_pred)
 apply (frule less_Suc_le1[of "i div s" "r - Suc 0"])
 apply (frule le_less_trans[of "Suc (i div s)" "r - Suc 0" "r"], assumption+)
 apply (simp only:mod_less) 
done

lemma Suc_rtos_div_r_3:"[|0 < r; 0 < s; i mod s < s - Suc 0; i ≤ r * s - Suc 0; 
      Suc (rtos r s i) < r * s; i div s < r - Suc 0|] ==> 
               Suc (rtos r s i) div r = i mod s"
apply (simp add:rtos_def) 
apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"]) 
apply simp
apply simp
apply (subst add_Suc_right[THEN sym, of "i mod s * r" "i div s"])
 apply (subst add_commute[of "i mod s * r" "Suc (i div s)"])
 apply (frule less_Suc_le1[of "i div s" "r - Suc 0"])
 apply (frule le_less_trans[of "Suc (i div s)" "r - Suc 0" "r"], simp)
 apply (subst div_mult_self1[of "r" "Suc (i div s)" "i mod s"], assumption)
 apply (subst div_less[of "Suc (i div s)" "r"], assumption)
 apply (simp only:add_0_right) 
done

lemma r_s_div_s:"[|0 < r; 0 < s|] ==> (r * s - Suc 0) div s = r - Suc 0"
apply (frule rfn_tool1_1 [of "s" "s - Suc 0" "r - Suc 0"]) 
 apply simp
 apply (simp add:diff_mult_distrib) 
done
 
lemma r_s_mod_s:"[|0 < r; 0 < s|] ==> (r * s - Suc 0) mod s = s - Suc 0"
apply (frule rfn_tool20 [of "s" "(r - Suc 0) * s + (s - Suc 0)" "r - Suc 0"
        "s - Suc 0"], simp, simp)
apply (simp add:diff_mult_distrib)
done

lemma rtos_r_s:"[|0 < r; 0 < s|] ==> rtos r s (r * s - Suc 0) = r * s - Suc 0"
apply (simp add:rtos_def)
apply (frule r_s_div_s [of "r" "s"], assumption+)
apply (frule r_s_mod_s [of "r" "s"], assumption+)
apply (simp, simp add:diff_mult_distrib, simp add:mult_commute)
done

lemma rtos_rs_1:"[| 0 < r; 0 < s; rtos r s i < r * s;
       ¬ Suc (rtos r s i) < r * s|] ==> rtos r s i = r * s - Suc 0"
apply (frule pos_mult_pos [of "r" "s"], assumption+)
apply (frule less_le_diff [of "rtos r s i" "r * s"])
apply (simp add:nat_not_le_less[THEN sym, of "Suc (rtos r s i)" "r * s"])
done

lemma rtos_rs_i_rs:"[| 0 < r; 0 < s; i ≤ r * s - Suc 0; 
rtos r s i = r * s - Suc 0|] ==>  i = r * s - Suc 0" 
apply (frule mem_of_Nset [of "i" "r * s - Suc 0"])     
apply (frule rtos_r_s [of "r" "s"], assumption+)
apply (frule rtos_inj[of "r" "s"], assumption+)
apply (cut_tac n_in_Nsetn[of "r * s - Suc 0"])
apply (simp add:inj_on_def)
done

lemma JHS_Tr1_1:"[|Group G; 0 < r; 0 < s; compseries G r f; compseries G s g|] ==> f (Suc ((r * s - Suc 0) div s)) \<struct>G (f ((r * s - Suc 0) div s) ∩ g ((r * s - Suc 0) mod s)) = f (r - Suc 0) ∩ g (s - Suc 0)"
apply (frule r_s_div_s [of "r" "s"], assumption+)
apply (frule r_s_mod_s [of "r" "s"], assumption+)
apply simp 
apply (subst Group.compseriesTr3 [of "G" "r" "f"], assumption+)
apply (frule Group.compseriesTr0 [of "G" "r" "f" "r - Suc 0"], assumption+)
 apply (simp add:less_mem_of_Nset)
apply (frule Group.compseriesTr0 [of "G" "s" "g" "s - Suc 0"], assumption+)
 apply (simp add:less_mem_of_Nset)
apply (frule Group.inter_sgs [of "G" "f (r - (Suc 0))" "g (s - Suc 0)"], 
                                                          assumption+)
apply (rule Group.s_top_l_unit, assumption+)
done

lemma JHS_Tr1_2:"[|Group G; 0 < r; 0 < s; compseries G r f; compseries G s g;
 k < r - Suc 0|] ==> ((Gp G (f (Suc k) \<struct>G (f k ∩ g (s - Suc 0)))) / 
                               (f (Suc (Suc k)) \<struct>G (f (Suc k) ∩ g 0))) ≅
               ((Gp G (g s \<struct>G (g (s - Suc 0) ∩ f k))) /
                          (g s \<struct>G (g (s - Suc 0) ∩ f (Suc k))))"
apply (frule Group.compseriesTr0 [of "G" "r" "f" "k"], assumption+)
apply (frule less_trans [of "k" "r - Suc 0" "r"]) apply simp
apply simp 
apply (frule Suc_leI [of "k" "r - Suc 0"]) 
apply (frule le_less_trans [of "Suc k" "r - Suc 0" "r"]) apply simp
apply (frule Group.compseriesTr0 [of "G" "r" "f" "Suc k"], assumption+)
apply simp
 apply (thin_tac "Suc k ≤ r - Suc 0")
 apply (frule Suc_leI[of "Suc k" "r"])
apply (frule Group.compseriesTr0 [of "G" "r" "f" "Suc (Suc k)"], assumption+)
apply (frule Group.compseriesTr0 [of "G" "s" "g" "s - Suc 0"], assumption+)
apply simp
apply (subst Group.compseriesTr2 [of "G" "s" "g"], assumption+)
apply (subst Group.compseriesTr3 [of "G" "s" "g"], assumption+)
apply (subst Int_absorb2 [of "f (Suc k)" "carrier G"])
 apply (rule Group.sg_subset, assumption+) 
 apply (subst Group.K_absorb_HK[of "G" "f (Suc (Suc k))" "f (Suc k)"], assumption+) 
 apply (rule Group.compseriesTr5 [of "G" "r" "f" "Suc k"], assumption+)
 apply (frule Suc_leI [of "k" "r - Suc 0"])
 apply simp
apply (subst Group.s_top_l_unit[of "G" "g (s - Suc 0) ∩ f k"], assumption+)
 apply (simp add:Group.inter_sgs)
apply (subst Group.compseriesTr3[of "G" "s" "g"], assumption+)
apply (subst Group.s_top_l_unit[of "G" "g (s - Suc 0) ∩ f (Suc k)"], assumption+)
 apply (simp add:Group.inter_sgs)
apply (frule Group.Group_Gp [of "G" "f k"], assumption+) 
apply (frule Group.compser_nsg [of "G" "r" "f" "k"], assumption+)
apply (simp add:less_mem_of_Nset [of "k" "r - Suc 0"])
apply (frule isom4b [of "Gp G (f k)" "f (Suc k)" "f k ∩ g (s - Suc 0)"], 
                                                          assumption+)
apply (rule Group.sg_sg, assumption+)
 apply (rule Group.inter_sgs, assumption+)
 apply (simp add:Int_lower1)
 apply (frule Group.compseriesTr5[of "G" "r" "f" "k"], assumption+)
 apply simp
 apply (frule Group.s_top_induced[of "G" "f k" "f (Suc k)" "f k ∩ g (s - Suc 0)"], assumption+)
 apply (simp add:Int_lower1) apply simp 
  apply (thin_tac "f (Suc k) \<struct>(\<natural>G(f k)) (f k ∩ g (s - Suc 0)) =
     f (Suc k) \<struct>G (f k ∩ g (s - Suc 0))")
  apply (frule Suc_pos [of "Suc k" "r"])
 apply (frule Group.cmp_rfn0 [of "G" "r" "s" "f" "g" "k" "s - Suc 0"], assumption+) 
 apply simp 
 apply simp
 apply (frule Group.sg_inc_set_mult[of "G" "f k" "f (Suc k)" "f k ∩ g (s - Suc 0)"], assumption+) apply (simp add:Int_lower1)

 apply (simp add:Group.Gp_inherited [of "G" "f (Suc k) \<struct>G (f k ∩ g (s - Suc 0))" "f k"])
  apply (frule Group.inter_sgs [of "G" "f k" "g (s - Suc 0)"], assumption+)
apply (frule Group.Gp_inherited [of "G" "f k ∩ g (s - Suc 0)" "f k"], assumption+)
apply (rule Int_lower1) apply simp
 apply (thin_tac "(Gp (Gp G (f k)) ((f k) ∩ (g (s - Suc 0)))) =
                                        (Gp G ((f k) ∩ (g (s - Suc 0))))")
 apply (thin_tac "f (Suc k) \<struct>G f k ∩ g (s - Suc 0) ⊆ f k")
 apply (thin_tac "G » f (Suc k) \<struct>G f k ∩ g (s - Suc 0)")
 apply (simp add:Int_assoc [of "f k" "g (s - Suc 0)" "f (Suc k)"])
 apply (simp add:Int_commute [of "g (s - Suc 0)" "f (Suc k)"])
 apply (simp add:Int_assoc [of "f k" "f (Suc k)" "g (s - Suc 0)", THEN sym])
 apply (simp add:Int_absorb1[of "f (Suc k)" "f k"])
apply (simp add:Int_commute)
done

lemma JHS_Tr1_3:"[|Group G; 0 < r; 0 < s; compseries G r f; compseries G s g;
       i ≤ s * r - Suc 0; Suc (rtos r s i) < s * r; Suc i < s * r;
       i mod s < s - Suc 0; Suc i div s ≤ r - Suc 0; i div s = r - Suc 0|]
    ==> Group (Gp G (f r \<struct>G (f (r - Suc 0) ∩ g (i mod s))) /
        (f r \<struct>G (f (r - Suc 0) ∩ g (Suc (i mod s)))))"
apply (frule nat_eq_le[of "i div s" "r - Suc 0"])
apply (frule Group.compser_nsg [of "G" "r" "f" "i div s"], assumption+)
 apply simp
apply (frule Group.compser_nsg [of "G" "s" "g" "i mod s"], assumption+)
 apply simp
 apply (frule  Group.compseriesTr0 [of "G" "r" "f" "r - Suc 0"], assumption+)
 apply simp 
 apply (frule  Group.compseriesTr0 [of "G" "r" "f" "r"], assumption+)
 apply simp 
 apply (frule  Group.compseriesTr0 [of "G" "s" "g" "i mod s"], assumption+)
 apply (simp add:less_imp_le)
 apply (frule  Group.compseriesTr0 [of "G" "s" "g" "Suc (i mod s)"], assumption+)
 apply (frule Suc_mono [of "i mod s" "s - (Suc 0)"],
        simp add:less_mem_of_Nset)
apply (frule Group.cmp_rfn0 [of "G" "r" "s" "f" "g" "i div s" "i mod s"], 
       assumption+, simp, simp) 
apply (frule  Group.ZassenhausTr3 [of "G" "f (r - Suc 0)" "f r" "g (i mod s)" 
       "g (Suc (i mod s))"], assumption+, simp, simp)
apply (cut_tac Group.Group_Gp [of "G" "f r \<struct>G (f (r - Suc 0) ∩ g (i mod s))"])
apply (rule Group.Group_Qg, assumption+)
apply simp
done

lemma JHS_Tr1_4:"[|Group G; 0 < r; 0 < s; compseries G r f; compseries G s g;
       i ≤ s * r - Suc 0; Suc (rtos r s i) < s * r; Suc i < s * r;
       i mod s < s - Suc 0; Suc i div s ≤ r - Suc 0; i div s = r - Suc 0|] ==> 
      Group (Gp G (g (Suc (i mod s)) \<struct>G (g (i mod s) ∩ f (r - Suc 0))) /
       (g (Suc (Suc (i mod s))) \<struct>G (g (Suc (i mod s)) ∩ f 0)))"
apply (subst Group.compseriesTr2 [of "G" "r" "f"], assumption+)
apply (frule Group.compseriesTr0 [of "G" "s" "g" "Suc (i mod s)"], assumption+)
apply (frule Suc_mono [of "i mod s" "s - Suc 0"], simp)
apply (frule Group.sg_subset [of "G" "g (Suc (i mod s))"], assumption+)
apply (subst Int_absorb2, assumption+)
apply (frule Suc_mono [of "i mod s" "s - Suc 0"])
apply (frule less_mem_of_Nset [of "i mod s" "s - Suc 0"], simp)
apply (frule Suc_leI [of "Suc (i mod s)" "s"])
apply (frule Group.compseriesTr0 [of "G" "s" "g" "Suc (Suc (i mod s))"], 
             assumption+)
 apply (frule less_le_diff [of "Suc (i mod s)" "s"])
 apply (frule Suc_pos [of "Suc (i mod s)" "s"])
 apply (frule Group.compseriesTr5[of "G" "s" "g" "Suc (i mod s)"], assumption+)
 apply (subst Group.K_absorb_HK[of "G" "g (Suc (Suc (i mod s)))"
                                        "g (Suc (i mod s))"], assumption+)
apply (frule Group.compseriesTr0 [of "G" "s" "g" "i mod s"], assumption+)
 apply (frule mod_less_divisor [of "s" "i"], simp)
 apply (frule Group.cmp_rfn0 [of "G" "s" "r" "g" "f" "i mod s" "r - Suc 0"], 
        assumption+, simp, simp, simp)
 apply (cut_tac Group.compser_nsg [of "G" "s" "g" "i mod s"])
apply (frule Group.ZassenhausTr4_1 [of "G" "g (i mod s)" "g (Suc (i mod s))" 
                    "f (r - Suc 0)"], assumption+)
 apply (rule Group.sg_sg, assumption+)
 apply (rule Group.inter_sgs, assumption+)
 apply (rule Group.compseriesTr0 [of "G" "r" "f" "r - Suc 0"], assumption+)
 apply simp 
 apply (rule Int_lower1)
apply (rule Group.Group_Qg)
 apply (rule Group.Group_Gp, assumption+, simp+)
done

lemma JHS_Tr1_5:"[|Group G; 0 < r; 0 < s; compseries G r f; compseries G s g;
      i ≤ s * r - Suc 0; Suc (rtos r s i) < s * r; Suc i < s * r;
      i mod s < s - Suc 0; i div s < r - Suc 0|]
 ==> (Gp G (f (Suc (i div s)) \<struct>G (f (i div s) ∩ g (i mod s))) /
        (f (Suc (i div s)) \<struct>G (f (i div s) ∩ g (Suc (i mod s))))) ≅
     (Gp G (g (Suc (i mod s)) \<struct>G (g (i mod s) ∩ f (i div s))) /
       (g (Suc (Suc (rtos r s i) div r)) \<struct>G 
              (g (Suc (rtos r s i) div r) ∩ f (Suc (rtos r s i) mod r))))"
apply (frule Group.compseriesTr0 [of "G" "s" "g" "i mod s"], assumption+)
 apply simp 
apply (frule Group.compseriesTr0 [of "G" "s" "g" "Suc (i mod s)"], assumption+)
apply (frule Suc_mono [of "i mod s" "s - Suc 0"], simp)
apply (frule Group.compseriesTr0 [of "G" "r" "f" "i div s"], assumption+)
 apply (frule less_trans [of "i div s" "r - Suc 0" "r"], simp)
 apply simp
apply (frule Group.compseriesTr0 [of "G" "r" "f" "Suc (i div s)"], assumption+)
 apply (frule Suc_mono [of "i div s" "r - Suc 0"]) 
 apply simp 
 apply (frule Group.compser_nsg [of "G" "r" "f" "i div s"], assumption+)
 apply simp 
 apply (frule Group.compser_nsg [of "G" "s" "g" "i mod s"], assumption+, simp)
apply (subst Suc_rtos_i_mod_r_3 [of "r" "s" "i"], assumption+)
 apply (simp add:mult_commute, assumption)
 apply (subst Suc_rtos_div_r_3 [of "r" "s" "i" ], assumption+)+
 apply (simp add:mult_commute)+
 apply (subst Suc_rtos_div_r_3[of "r" "s" "i"], assumption+)
apply (rule Zassenhaus [of "G" "f (i div s)" "f (Suc (i div s))" "g (i mod s)"
 "g (Suc (i mod s))"], assumption+)
done 

lemma JHS_Tr1_6:" [|Group G; 0 < r; 0 < s; compseries G r f; compseries G s g;
 i ≤ r * s - Suc 0; Suc (rtos r s i) < r * s|] ==>
  ((Gp G (cmp_rfn G r f s g i)) / (cmp_rfn G r f s g (Suc i))) ≅
  ((Gp G (g (Suc (rtos r s i div r)) \<struct>G 
      (g (rtos r s i div r) ∩ f (rtos r s i mod r)))) /
      (g (Suc (Suc (rtos r s i) div r)) \<struct>G
              (g (Suc (rtos r s i) div r) ∩ f (Suc (rtos r s i) mod r))))"
apply (simp add:cmp_rfn_def)
 apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp)
 apply (simp add:mult_commute [of "r" "s"])
 apply (frule Suc_leI [of "i" "s * r"], thin_tac "i < s * r")
apply (case_tac "¬ Suc i < s * r", simp)

 apply (frule rfn_tool17 [of "Suc i" "s * r" "Suc 0"])
 apply (thin_tac " Suc i = s * r")
 apply simp
 apply (frule rtos_r_s [of "r" "s"], assumption+) 
 apply (simp add:mult_commute [of "r" "s"])  (* !! ??? *) 
apply simp
apply (frule mod_less_divisor [of "s" "i"])
apply (frule less_le_diff [of "i mod s" "s"], thin_tac "i mod s < s")
 apply (case_tac "i mod s = s - Suc 0", simp)
 apply (frule_tac div_Tr2 [of "r" "s" "Suc i"], assumption+)
 apply (simp add:le_imp_less_or_eq)
 apply (subst div_Tr3_1[of "r" "s" "i"], assumption+, simp)
 apply (subst rtos_hom3 [of "r" "s" "i"], assumption+) 
 apply (simp add:mult_commute)
 apply (subst rtos_hom3_1 [of "r" "s" "i"], assumption+)
 apply (simp add:mult_commute)
apply (frule div_Tr3_1 [of "r" "s" "i"], assumption+, simp)
 apply (simp, thin_tac "Suc i div s = Suc (i div s)")
 apply (insert n_less_Suc [of "i div s"])
 apply (frule less_le_trans [of "i div s" "Suc (i div s)" "r - Suc 0"], 
                                                             assumption+)
 apply (subst Suc_rtos_div_r_1 [of "r" "s" "i"], assumption+) 
 apply (simp add:mult_commute[of "s" "r"])+ 
 apply (subst Suc_rtos_mod_r_1 [of "r" "s" "i"], assumption+)
 apply (subst Suc_i_mod_s_0_1 [of "r" "s" "i"], assumption+)
 apply (simp only:Suc_rtos_div_r_1 [of "r" "s" "i"])  
 apply (subst rtos_hom3[of "r" "s" "i"], assumption+, simp)

 apply (frule JHS_Tr1_2 [of "G" "r" "s" "f" "g" "i div s"], assumption+,
        simp)
apply (frule noteq_le_less [of "i mod s" "s - Suc 0"], assumption+)
 apply (thin_tac "i mod s ≤ s - Suc 0")
 apply (thin_tac "i mod s ≠ s - Suc 0")
 apply (frule div_Tr2 [of "r" "s" "Suc i"], assumption+,
        rule noteq_le_less, assumption+)
 apply (subst div_Tr3_2 [THEN sym, of "r" "s" "i"], assumption+)
 apply simp
 apply (subst rfn_tool12_1 [THEN sym, of "s" "i"], assumption+)
 apply simp
 apply (subst rtos_hom3 [of "r" "s" "i"], assumption+)
 apply (simp add:mult_commute)
 apply (subst rtos_hom3_1 [of "r" "s" "i"], assumption+)
 apply (simp add:mult_commute)
apply (case_tac "i div s = r - Suc 0")
apply (subst rtos_hom5 [of "r" "s" "i"], assumption+)
 apply (simp add:mult_commute)
 apply assumption
 apply (subst rtos_hom7 [of "r" "s" "i"], assumption+)
 apply (simp add:mult_commute)
 apply (assumption, simp)
apply (frule JHS_Tr1_3 [of "G" "r" "s" "f" "g" "i"], assumption+,
       simp only:noteq_le_less,  assumption+)
apply (frule JHS_Tr1_4 [of "G" "r" "s" "f" "g" "i"], assumption+,
       simp only:noteq_le_less, assumption+)
 apply (subst rtos_hom3 [of "r" "s" "i"], assumption+,
        simp only:mult_commute[of "s" "r"])
 apply (subst rtos_hom5 [of "r" "s" "i"], assumption+,
        simp only:mult_commute[of "s" "r"],
               simp add:mem_of_Nset) 
apply (rule isomTr1, assumption+)
 apply (frule Suci_div_s_2[of "r" "s" "i"], assumption+,
        simp only:mult_commute, assumption)
 apply simp 
apply (frule Suci_div_s_2[of "r" "s" "i"], assumption+,
       simp only:mult_commute, assumption, simp)
apply (rule JHS_Tr1_2 [of "G" "s" "r" "g" "f" "i mod s"], assumption+)
apply (frule div_Tr2 [of "r" "s" "i"], assumption+)
 apply (rule le_less_trans [of "i" "s * r - Suc 0" "s * r"], assumption+)
 apply simp
apply (frule noteq_le_less [of "i div s" "r - Suc 0"], assumption+)
 apply (frule Suci_div_s_2[of "r" "s" "i"], assumption+,
        simp only:mult_commute[of "s" "r"], assumption, simp)
apply (frule JHS_Tr1_5[of "G" "r" "s" "f" "g" "i"], assumption+)
 apply (simp add:noteq_le_less[of "Suc i"], assumption+)
 apply (frule mem_of_Nset[of "i" "s*r - Suc 0"], 
                        simp add:mult_commute[of "s" "r"])
 apply (simp add:rtos_hom3 [of "r" "s" "i"])
done 

lemma JHS_Tr1:"[| Group G; 0 < r; 0 < s; compseries G r f; compseries G s g|]
==> isom_Gchains (r * s - 1) (rtos r s) (Qw_cmpser G (cmp_rfn G r f s g)) (Qw_cmpser G (cmp_rfn G s g r f))"
apply (simp add:isom_Gchains_def) 
 apply (rule allI, rule impI) 
 apply (frule pos_mult_pos [of "r" "s"], assumption+)
 apply (frule_tac b = "r * s" and a = i in rfn_tool11, assumption+)
apply (simp add:Qw_cmpser_def)
 apply (simp only:cmp_rfn_def [of "G" "s" "g"]) 
 apply (frule_tac l = i in rtos_hom1 [of "r" "s"], assumption+)
 apply (frule_tac x = "rtos r s i" and y = "s * r - Suc 0" and z = "s * r" in
          le_less_trans, simp)
 apply (simp add:mult_commute [of "s" "r"])
apply (case_tac "Suc (rtos r s i) < r * s", simp)
prefer 2 apply simp
 apply (frule_tac i = i in rtos_rs_1 [of "r" "s"], assumption+) 
 apply (frule_tac i = i in rtos_rs_i_rs [of "r" "s"], assumption+)
 apply (rule less_le_diff, assumption+)
 apply (simp add:cmp_rfn_def)
 apply (simp add:mult_commute)
apply (subst JHS_Tr1_1 [of "G" "r" "s" "f" "g"], assumption+)
apply (frule JHS_Tr1_1 [of "G" "s" "r" "g" "f"], assumption+)
 apply (simp add:mult_commute [of "r" "s"])
 apply (simp add:Int_commute) 
apply (frule Group.compseriesTr0 [of "G" "r" "f" "r - Suc 0"], assumption+,
       simp)
apply (frule Group.compseriesTr0 [of "G" "s" "g" "s - Suc 0"], assumption+,
       simp)
apply (frule Group.inter_sgs [of "G" "f (r - (Suc 0))" "g (s - Suc 0)"], 
                                                          assumption+)
apply (frule Group.special_sg_e [of "G"])
apply (frule Group.special_nsg_e [of "G" "f (r - Suc 0) ∩ g (s - Suc 0)"], 
                                                        assumption+)
apply (frule Group.Group_Gp [of "G" "f (r - Suc 0) ∩ g (s - Suc 0)"], assumption+)
apply (frule Group.Group_Qg[of "Gp G (f (r - Suc 0) ∩ g (s - Suc 0))" "{\<one>G}"],
       assumption+)
apply (simp add:isomTr0[of "(\<natural>G(f (r - Suc 0) ∩ g (s - Suc 0))) / {\<one>G}"])
apply (rule JHS_Tr1_6, assumption+)
done
 
lemma abc_SucTr0:"[|(0::nat) < a; c ≤ b; a - Suc 0 = b - c|] ==> a = (Suc b) - c"
apply (subgoal_tac "Suc 0 ≤ a")
apply (frule le_add_diff_inverse2 [of "Suc 0" "a", THEN sym])
apply auto 
done

lemma length_wcmpser0_0:"[|Group G; Ugp E; w_cmpser G (Suc 0) f|] ==> 
       f ` {i. i ≤ (Suc 0)} = {f 0, f (Suc 0)}"
apply (simp add:Nset_1)
done

lemma length_wcmpser0_1:"[|Group G; Ugp E; w_cmpser G (Suc n) f; i∈{i. i ≤ n};
 (Qw_cmpser G f) i ≅ E|] ==> f i = f (Suc i)"
apply (subgoal_tac "0 < Suc n")
apply (frule Group.w_cmpser_gr [of "G" "Suc n" "f" "i"], assumption+) 
prefer 2 apply simp
apply (frule Group.w_cmpser_gr [of "G" "Suc n" "f" "Suc i"], simp+) 
apply (frule Group.w_cmpser_ns [of "G" "Suc n" "f" "i"], simp+)
apply (simp add:Qw_cmpser_def)
apply (rule QgrpUnit_3 [of "G" "E" "f i" "f (Suc i)"], assumption+, simp+)
done

lemma length_wcmpser0_2:"[|Group G; Ugp E; w_cmpser G (Suc n) f; i ≤ n;
 ¬ (Qw_cmpser G f) i ≅ E|] ==> f i ≠ f (Suc i)"       
apply (cut_tac zero_less_Suc[of "n"])
apply (frule Group.w_cmpser_gr [of "G" "Suc n" "f" "i"], assumption+) 
apply simp
apply (frule Group.w_cmpser_gr [of "G" "Suc n" "f" "Suc i"], assumption+) 
apply simp
apply (frule Group.w_cmpser_ns [of "G" "Suc n" "f" "i"], assumption+, simp)
apply (simp add:Qw_cmpser_def)
apply (rule QgrpUnit_4 [of "G" "E" "f i" "f (Suc i)"], assumption+)
done

lemma length_wcmpser0_3:"[|Group G; Ugp E; w_cmpser G (Suc (Suc n)) f;
  f (Suc n) ≠ f (Suc (Suc n))|] ==> f (Suc (Suc n)) ∉ f ` {i. i ≤ (Suc n)}"
apply (frule Group.w_cmpser_is_d_gchain, assumption+)
apply (thin_tac "w_cmpser G (Suc (Suc n)) f")
apply (rule contrapos_pp, simp+)
apply (frule Group.d_gchainTr2 [of "G" "Suc ((Suc n))" "f" "Suc n" "Suc (Suc n)"])
apply simp apply assumption+ apply simp+
apply (frule psubsetI [of "f (Suc (Suc n))" "f (Suc n)"])
apply (rule not_sym, assumption+) 
apply (thin_tac "f (Suc (Suc n)) ⊆ f (Suc n)")
apply (simp add:image_def)
apply (erule exE)
apply (cut_tac zero_less_Suc[of "Suc n"])
apply (frule_tac f = f and l = x in Group.d_gchainTr2 [of "G" "Suc ((Suc n))" 
                              _ _ "Suc (Suc n)"], assumption+)
 apply simp+
apply (frule_tac f = f and l = x in Group.d_gchainTr2 [of "G" "Suc (Suc n)" 
                              _ _ "Suc n"], simp+)
done

lemma length_wcmpser0_4:"[|Group G; Ugp E; w_cmpser G (Suc 0) f|] ==>
  card (f ` {i. i ≤ Suc 0}) - 1 = Suc 0 - card {i. i = 0 ∧ 
                                                Qw_cmpser G f i ≅ E}"
 (* card (f ` Nset (Suc 0)) - 1 =
           Suc 0 - card {i. i ∈ Nset 0 ∧ Qw_cmpser G f i ≅ E}" *)
apply (subst length_wcmpser0_0, assumption+)
apply (case_tac "Qw_cmpser G f 0 ≅ E")
 apply (frule_tac n = 0 and f = f and i = 0 in length_wcmpser0_1 [of "G" "E"],         assumption+, simp+)
 apply (rule nonempty_card_pos1[of "{i. i = 0 ∧ Qw_cmpser G f i ≅ E}"])
 apply (cut_tac finite1[of "(0::nat)"])
 apply (rule finite_subset[of "{i. i = 0 ∧ Qw_cmpser G f i ≅ E}" "{0}"])
  apply (rule subsetI) apply simp apply assumption
  apply blast

 apply (frule_tac f = f and i = 0 in length_wcmpser0_2 [of "G" "E" "0"], 
                                                        assumption+)
 apply (simp add:Nset_inc_0, assumption)
 apply (cut_tac finite1[of "f (Suc 0)"])
 apply (simp only:singleton_iff[THEN sym, of "f 0" "f (Suc 0)"])
 apply (simp add: card_insert_disjoint [of "{f (Suc 0)}" "f 0"])
 apply (rule diff_zero_eq[of "card {i. i = 0 ∧ Qw_cmpser G f i ≅ E}" 
        "Suc 0"]) 
 apply (rule sym)
 apply (subst card0[THEN sym])
 apply (rule_tac A = "{}" and B = "{i. i = 0 ∧ Qw_cmpser G f i ≅ E}" in 
        card_eq) 
 apply (rule equalityI, rule subsetI) 
 apply simp
apply (rule subsetI, simp, erule conjE, simp)
done

lemma length_wcmpser0_5:" [|Group G; Ugp E; w_cmpser G (Suc (Suc n)) f; 
      w_cmpser G (Suc n) f;  
        card (f ` {i. i ≤ (Suc n)}) - 1 = Suc n - 
                     card {i. i ≤ n  ∧ Qw_cmpser G f i ≅ E}; 
        Qw_cmpser G f (Suc n) ≅ E|] ==>
      card (f ` {i . i ≤ (Suc (Suc n))}) - 1 =
             Suc (Suc n) - card {i. i ≤ (Suc n) ∧ Qw_cmpser G f i ≅ E}"
apply (frule_tac n = "Suc n" and f = f and i = "Suc n" in 
                            length_wcmpser0_1 [of "G" "E"], assumption+)
 apply (simp, assumption)
apply (subgoal_tac "f ` {i. i ≤ (Suc (Suc n))} = f ` {i. i ≤ (Suc n)}")
 apply simp
  prefer 2  apply (rule equalityI) 
  apply (simp add:image_def)
  apply (auto del:equalityI)
  apply (case_tac "xa = Suc (Suc n)", simp)  
  apply (thin_tac " xa = Suc (Suc n)", rotate_tac -2)
  apply (frule sym, thin_tac "f (Suc n) = f (Suc (Suc n))",
         simp, thin_tac "f (Suc (Suc n)) = f (Suc n)")
 apply (cut_tac n_in_Nsetn[of "Suc n"], blast)
  apply (frule_tac m = xa and n = "Suc (Suc n)" in noteq_le_less, assumption,
         frule_tac x = xa in Suc_less_le[of _ "Suc n"], blast) 
apply (subgoal_tac "{i. i ≤ (Suc n) ∧ Qw_cmpser G f i ≅ E} =
              insert (Suc n) {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}")
apply simp
apply (subgoal_tac "(Suc n) ∉ {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}")
apply (subgoal_tac "finite {i. i ≤ n ∧ Qw_cmpser G  f i ≅ E}")
apply (subst card_insert_disjoint)
 apply (assumption+, simp)
 apply (cut_tac finite_Nset [of "n"])
 apply (subgoal_tac "{i. i ≤ n ∧ Qw_cmpser G f i ≅ E} ⊆ {i. i ≤ n}")

apply (rule finite_subset, assumption+)
apply (rule subsetI, simp add:CollectI)
apply simp
apply (rule equalityI)
 apply (rule subsetI, simp add:CollectI) 
 apply (thin_tac "card (f ` {i. i ≤ (Suc n)}) - Suc 0 =
          Suc n - card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}")
 apply (thin_tac "f ` {i. i ≤ (Suc (Suc n))} = f ` {i. i ≤ (Suc n)}")
 apply (erule conjE) 
 apply (case_tac "x = Suc n", simp) 
 apply (frule_tac m = x and n = "Suc n" in noteq_le_less, assumption+)
 apply (frule_tac x = x and n = "Suc n" in less_le_diff, simp) 
 apply (rule subsetI) 
 apply (thin_tac "card (f ` {i. i ≤ (Suc n)}) - Suc 0 =
          Suc n - card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}")
 apply (thin_tac "f ` {i. i ≤ (Suc (Suc n))} = f ` {i. i ≤ (Suc n)}")
 apply simp
 apply (case_tac "x = Suc n") apply simp+
done

lemma length_wcmpser0_6:"[|Group G; w_cmpser G (Suc (Suc n)) f|] ==> 
                                          0 < card (f ` {i. i ≤ (Suc n)})"
apply (insert finite_Nset [of "Suc n"])
apply (frule finite_imageI [of "{i. i ≤ (Suc n)}" "f"])
apply (subgoal_tac "{f 0} ⊆ f ` {i. i ≤ (Suc n)}")
apply (frule card_mono [of "f ` {i. i ≤ (Suc n)}" "{f 0}"], assumption+)
apply (simp add:card1 [of "f 0"])
apply (rule subsetI, simp add:image_def)
apply (subgoal_tac "0 ∈ {i. i ≤ (Suc n)}", blast)
apply simp
done

lemma length_wcmpser0_7:"[|Group G; w_cmpser G (Suc (Suc n)) f|] ==>
                     card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E} ≤ Suc n"
apply (insert finite_Nset [of "n"])
apply (subgoal_tac "{i. i ≤ n ∧ Qw_cmpser G f i ≅ E} ⊆ {i. i ≤ n}")
apply (frule card_mono [of "{i. i ≤ n}" "{i. i ≤ n ∧ Qw_cmpser G f i ≅ E}"])
 apply (assumption, simp add:card_Nset)
apply (rule subsetI, simp add:CollectI)
done


lemma length_wcmpser0:"[|Group G; Ugp E|] ==>∀f. w_cmpser G (Suc n) f -->  
card (f ` {i. i ≤ (Suc n)}) - 1 = (Suc  n) - (card {i. i ≤ n ∧ 
                                      ((Qw_cmpser G f) i ≅ E)})"
apply (induct_tac n)  
 (*** n = 0 ***)
apply (rule allI) apply (rule impI)
apply (frule_tac f = f in length_wcmpser0_4[of G E], assumption+, simp)
  (**** n ****)
 apply (rule allI) apply (rule impI)
 apply (frule_tac n = "Suc n" and f = f in Group.w_cmpser_pre [of "G"], assumption+)
 apply (subgoal_tac "card (f ` {i. i ≤ (Suc n)}) - 1 =
                 Suc n - card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}")
 prefer 2 apply simp
 apply (thin_tac " ∀f. w_cmpser G (Suc n) f -->
                 card (f ` {i. i ≤ (Suc n)}) - 1 =
                 Suc n - card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}")
apply (case_tac "Qw_cmpser G f (Suc n) ≅ E")  
apply (rule length_wcmpser0_5, assumption+)
apply (frule_tac n = "Suc n" and f = f and i = "Suc n" in 
                            length_wcmpser0_2 [of "G" "E"], assumption+)
 apply simp  apply assumption
 apply (subgoal_tac "f ` {i. i ≤ (Suc (Suc n))} = 
                       insert (f (Suc (Suc n))) (f ` {i. i ≤ (Suc n)})")
 apply simp apply (thin_tac "f ` {i. i ≤ (Suc (Suc n))} =
             insert (f (Suc (Suc n))) (f ` {i. i ≤ (Suc n)})")
 apply (subgoal_tac "finite (f ` {i. i ≤ (Suc n)})")
 apply (subgoal_tac "f (Suc (Suc n)) ∉ f ` {i. i ≤ (Suc n)}")
 apply (subst card_insert_disjoint, assumption) 
 apply assumption
 prefer 2  apply (rule length_wcmpser0_3, assumption+)
 prefer 2
  apply (subgoal_tac "finite {i. i ≤ (Suc n)}")
  apply (rule finite_imageI, assumption+, simp add:finite_Nset)
 prefer 2 
 apply (thin_tac " ¬ Qw_cmpser G f (Suc n) ≅ E",
        thin_tac " w_cmpser G (Suc n) f",
        thin_tac "f (Suc n) ≠ f (Suc (Suc n))")
 apply (subgoal_tac "{i. i ≤ (Suc (Suc n))} = {i. i≤(Suc n)} ∪ {Suc (Suc n)}")
 prefer 2 apply (rule_tac n = "Suc n" in Nset_un, simp)
 apply (subgoal_tac "card {i. i ≤ (Suc n) ∧ Qw_cmpser G f i ≅ E} =
                        card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}")
 apply simp 
 apply (thin_tac " card {i. i ≤ (Suc n) ∧ Qw_cmpser G f i ≅ E} =
             card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}",
        thin_tac "¬ Qw_cmpser G f (Suc n) ≅ E",
        thin_tac "f (Suc n) ≠ f (Suc (Suc n))",
        thin_tac "finite (f ` {i. i ≤ (Suc n)})",
        thin_tac "f (Suc (Suc n)) ∉ f ` {i. i ≤ (Suc n)}")
 apply (frule_tac n = n and f = f in length_wcmpser0_6 [of "G"], assumption+,
        frule_tac n = n and f = f in length_wcmpser0_7 [of "G"], assumption+)
 apply (rule abc_SucTr0, assumption+)
apply (rule card_eq)
 apply (thin_tac "card (f ` {i. i ≤ (Suc n)}) - Suc 0 =
             Suc n - card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}",
        thin_tac "f (Suc n) ≠ f (Suc (Suc n))",
        thin_tac "f (Suc (Suc n)) ∉ f ` {i. i ≤ (Suc n)}")
 apply (rule equalityI)
 apply (rule subsetI, simp add:CollectI, erule conjE)
 apply (case_tac "x = Suc n", simp, simp)
 apply (rule subsetI, simp add:CollectI)
done


lemma length_of_twcmpser:"[|Group G; Ugp E; tw_cmpser G (Suc n) f |] ==> 
      length_twcmpser G (Suc n) f = 
                     (Suc n) - (card {i. i ≤ n ∧ ((Qw_cmpser G f) i ≅ E)})"
apply (unfold length_twcmpser_def)
apply (insert length_wcmpser0 [of "G" "E" "n"])
apply (subgoal_tac "w_cmpser G (Suc n) f", rotate_tac -1,
       simp, simp, 
       thin_tac "∀f. w_cmpser G (Suc n) f -->
           card (f ` {i. i ≤ (Suc n)}) - Suc 0 =
           Suc n - card {i. i ≤ n ∧ Qw_cmpser G f i ≅ E}")
 apply (simp add:tw_cmpser_def w_cmpser_def, erule conjE)
 apply (thin_tac "∀i≤ n. Gp G (f i) \<triangleright> f (Suc i)")
 apply (simp add:td_gchain_def)
done 

lemma JHS_1:"[|Group G; Ugp E; compseries G r f; compseries G s g; 0<r; 0 < s|]
 ==> r =  r * s - card {i. i ≤ (r * s - Suc 0) ∧
            Qw_cmpser G (cmp_rfn G r f s g) i ≅ E}"
apply (frule_tac r = r and s = s and G = G and f = f and g = g in 
                   Group.JHS_Tr0, assumption+)
apply (simp add:wcsr_rfns_def, erule conjE)
apply (frule_tac length_of_twcmpser [of "G" "E" "r * s - Suc 0" 
        "cmp_rfn G r f s g"], assumption+, simp add:mult_commute)
apply (simp add:length_twcmpser_def)
apply (frule Group.rfn_compseries_iM [of "G" "r" "s" "f" "cmp_rfn G r f s g"], assumption+, rule Group.JHS_Tr0 [of "G" "r" "s" "f" "g"], assumption+)
apply (simp add:mult_commute [of "s" "r"])
done

lemma J_H_S:"[|Group G; Ugp E; compseries G r f; compseries G s g; 0<r;
               (0::nat)<s |]  ==> r = s"
apply (frule JHS_1 [of "G" "E" "r" "f" "s" "g"], assumption+,
       frule JHS_1 [of "G" "E" "s" "g" "r" "f"], assumption+,
       frule JHS_Tr1 [of "G" "r" "s" "f" "g"], assumption+,
       frule Group.JHS_Tr0 [of "G" "r" "s" "f" "g"], assumption+,
       frule Group.JHS_Tr0 [of "G" "s" "r" "g" "f"], assumption+)
 apply (simp add:wcsr_rfns_def, (erule conjE)+,
       frule Group.tw_cmpser_is_w_cmpser [of "G" "s * r" "cmp_rfn G r f s g"],
                                                        assumption+,
       frule Qw_cmpser [of "G" "s * r - Suc 0" "cmp_rfn G r f s g"],
       simp add:pos_mult_pos [of "s" "r"])
 apply (frule Group.tw_cmpser_is_w_cmpser [of "G" "r * s" "cmp_rfn G s g r f"],
                                                        assumption+,
        frule Qw_cmpser [of "G" "r * s - Suc 0" "cmp_rfn G s g r f"],
        simp add:pos_mult_pos [of "r" "s"],
        simp add:mult_commute [of "s" "r"])
apply (frule isom_gch_units [of "E" "r * s - Suc 0" 
 "Qw_cmpser G (cmp_rfn G r f s g)" "Qw_cmpser G (cmp_rfn G s g r f)" 
                                               "rtos r s"], assumption+)
prefer 2 apply simp
apply (simp add:Gch_bridge_def)
apply (rule conjI) apply (rule allI, rule impI) 
 apply (frule_tac l = l in rtos_hom1 [of "r" "s"], assumption+,
        simp add:mult_commute [of "s" "r"])
apply (rule rtos_inj, assumption+)
done 

end

5. Setproducts

lemma contain_commutator:

  [| G » H ; commutators G  H |] ==> G \<triangleright> H

lemma s_top_induced:

  [| G » L ; H  L; K  L |] ==> H \<struct>\<natural>L K = H \<struct>G K

lemma s_top_l_unit:

  G » K  ==> {\<one>} \<struct>G K = K

lemma s_top_r_unit:

  G » K  ==> K \<struct>G {\<one>} = K

lemma s_top_sub:

  [| G » H ; G » K  |] ==> H \<struct>G K  carrier G

lemma sg_inc_set_mult:

  [| G » L ; H  L; K  L |] ==> H \<struct>G K  L

lemma s_top_sub1:

  [| H  carrier G; K  carrier G |] ==> H \<struct>G K  carrier G

lemma s_top_elem:

  [| G » H ; G » K ; aH; bK |] ==> Group.top G a bH \<struct>G K

lemma s_top_elem1:

  [| H  carrier G; K  carrier G; aH; bK |]
  ==> Group.top G a bH \<struct>G K

lemma mem_s_top:

  [| H  carrier G; K  carrier G; uH \<struct>G K |]
  ==> ∃aH. ∃bK. Group.top G a b = u

lemma s_top_mono:

  [| H  carrier G; K  carrier G; H1.0  H; K1.0  K |]
  ==> H1.0 \<struct>G K1.0  H \<struct>G K

lemma s_top_unit_closed:

  [| G » H ; G » K  |] ==> \<one> ∈ H \<struct>G K

lemma s_top_commute:

  [| G » H ; G » K ; K \<struct>G H = H \<struct>G K; uH \<struct>G K;
     vH \<struct>G K |]
  ==> Group.top G u vH \<struct>G K

lemma s_top_commute1:

  [| G » H ; G » K ; K \<struct>G H = H \<struct>G K; uH \<struct>G K |]
  ==> ρ uH \<struct>G K

lemma s_top_commute_sg:

  [| G » H ; G » K ; K \<struct>G H = H \<struct>G K |] ==> G » H \<struct>G K 

lemma s_top_assoc:

  [| G » H ; G » K ; G » L  |]
  ==> H \<struct>G K \<struct>G L = H \<struct>G (K \<struct>G L)

lemma s_topTr6:

  [| G » H1.0 ; G » H2.0 ; G » K ; H1.0  K |]
  ==> (H1.0 \<struct>G H2.0) ∩ K = H1.0 \<struct>G H2.0K

lemma s_topTr6_1:

  [| G » H1.0 ; G » H2.0 ; G » K ; H2.0  K |]
  ==> (H1.0 \<struct>G H2.0) ∩ K = H1.0K \<struct>G H2.0

lemma l_sub_smult:

  [| G » H ; G » K  |] ==> H  H \<struct>G K

lemma r_sub_smult:

  [| G » H ; G » K  |] ==> K  H \<struct>G K

lemma s_topTr8:

  G » H  ==> H = H \<struct>G H

6. preliminary lemmas for Zassenhaus

lemma Gp_sg_subset:

  [| G » H ; \<natural>H » K  |] ==> K  H

lemma inter_Gp_nsg:

  [| G \<triangleright> N; G » H  |] ==> \<natural>H \<triangleright> HN

lemma ZassenhausTr0:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> \<natural>HK \<triangleright> HK1.0

lemma lcs_sub_s_mult:

  [| G » H ; G » N ; aH |] ==> aN  H \<struct>G N

lemma rcs_sub_smult:

  [| G » H ; G » N ; aH |] ==> Na  N \<struct>G H

lemma smult_commute_sg_nsg:

  [| G » H ; G \<triangleright> N |] ==> H \<struct>G N = N \<struct>G H

lemma smult_sg_nsg:

  [| G » H ; G \<triangleright> N |] ==> G » H \<struct>G N 

lemma smult_nsg_sg:

  [| G » H ; G \<triangleright> N |] ==> G » N \<struct>G H 

lemma Gp_smult_sg_nsg:

  [| G » H ; G \<triangleright> N |] ==> Group (\<natural>H \<struct>G N)

lemma N_sg_HN:

  [| G » H ; G \<triangleright> N |] ==> \<natural>H \<struct>G N » N 

lemma K_absorb_HK:

  [| G » H ; G » K ; H  K |] ==> H \<struct>G K = K

lemma nsg_Gp_nsg:

  [| G » H ; G \<triangleright> N; N  H |] ==> \<natural>H \<triangleright> N

lemma Gp_smult_nsg:

  [| G » H ; G \<triangleright> N |]
  ==> \<natural>H \<struct>G N \<triangleright> N

lemma Gp_smult_nsg1:

  [| G » H ; G \<triangleright> N |]
  ==> \<natural>N \<struct>G H \<triangleright> N

lemma ZassenhausTr2_3:

  [| G » H ; G » H1.0 ; \<natural>H \<triangleright> H1.0 |] ==> H1.0  H

lemma ZassenhausTr2_4:

  [| G » H ; G » H1.0 ; \<natural>H \<triangleright> H1.0; hH; h1.0H1.0 |]
  ==> Group.top G (Group.top G h h1.0) (ρ h) ∈ H1.0

lemma ZassenhausTr1:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> H1.0 \<struct>G HK1.0 = HK1.0 \<struct>G H1.0

lemma ZassenhausTr1_1:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> G » H1.0 \<struct>G HK1.0 

lemma ZassenhausTr2:

  [| G » H ; G » H1.0 ; G » K ; \<natural>H \<triangleright> H1.0 |]
  ==> H1.0 \<struct>G HK = HK \<struct>G H1.0

lemma ZassenhausTr2_1:

  [| G » H ; G » H1.0 ; G » K ; \<natural>H \<triangleright> H1.0 |]
  ==> G » H1.0 \<struct>G HK 

lemma ZassenhausTr2_2:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> H1.0 \<struct>G HK1.0  H1.0 \<struct>G HK

lemma ZassenhausTr2_5:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0; aH1.0; bHK1.0; cH1.0 |]
  ==> Group.top G (Group.top G a b) cH1.0 \<struct>G HK1.0

lemma ZassenhausTr2_6:

  [| u ∈ carrier G; v ∈ carrier G; x ∈ carrier G; y ∈ carrier G |]
  ==> Group.top G (Group.top G (Group.top G u v) (Group.top G x y))
       (ρ Group.top G u v) =
      Group.top G
       (Group.top G (Group.top G (Group.top G (Group.top G u v) x) (ρ v))
         (Group.top G (Group.top G v y) (ρ v)))
       (ρ u)

lemma ZassenhausTr2_7:

  [| a ∈ carrier G; x ∈ carrier G; y ∈ carrier G |]
  ==> Group.top G (Group.top G a (Group.top G x y)) (ρ a) =
      Group.top G (Group.top G (Group.top G a x) (ρ a))
       (Group.top G (Group.top G a y) (ρ a))

lemma ZassenhausTr3:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> \<natural>H1.0 \<struct>G HK \<triangleright> H1.0 \<struct>G HK1.0

lemma ZassenhausTr3_2:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> G » H1.0 \<struct>G HK1.0 \<struct>G HK 

lemma ZassenhausTr3_3:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> H1.0K \<struct>G HK1.0 = K1.0H \<struct>G KH1.0

lemma ZassenhausTr3_4:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0; gHK; hHK1.0 |]
  ==> Group.top G (Group.top G g h) (ρ g) ∈ HK1.0

lemma ZassenhausTr3_5:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> \<natural>HK \<triangleright> H1.0K \<struct>G HK1.0

lemma ZassenhausTr4:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> H1.0 \<struct>G HK1.0 \<struct>G (H1.0 \<struct>G HK) =
      H1.0 \<struct>G HK

lemma ZassenhausTr4_0:

  [| G » H ; G » H1.0 ; G » K ; G » K1.0 ; \<natural>H \<triangleright> H1.0;
     \<natural>K \<triangleright> K1.0 |]
  ==> H1.0 \<struct>G HK = H1.0 \<struct>G HK1.0 \<struct>G HK

lemma ZassenhausTr4_1:

  [| G » H ; \<natural>H \<triangleright> H1.0; \<natural>H » HK  |]
  ==> \<natural>H1.0 \<struct>G HK \<triangleright> H1.0

7. homomorphism

lemma gHom:

  [| Group F; Group G; f ∈ gHom F G; x ∈ carrier F; y ∈ carrier F |]
  ==> f (Group.top F x y) = Group.top G (f x) (f y)

lemma gHom_mem:

  [| Group F; Group G; f ∈ gHom F G; x ∈ carrier F |] ==> f x ∈ carrier G

lemma gHom_func:

  [| Group F; Group G; f ∈ gHom F G |] ==> f ∈ carrier F -> carrier G

lemma gHomcomp:

  [| Group F; Group G; Group H; f ∈ gHom F G; g ∈ gHom G H |]
  ==> g oF f ∈ gHom F H

lemma gHom_comp_gsurjec:

  [| Group F; Group G; Group H; gsurjF,G f; gsurjG,H g |] ==> gsurjF,H (g oF f)

lemma gHom_comp_ginjec:

  [| Group F; Group G; Group H; ginjF,G f; ginjG,H g |] ==> ginjF,H (g oF f)

lemma ghom_unit_unit:

  [| Group F; Group G; f ∈ gHom F G |] ==> f \<one>F = \<one>G

lemma ghom_inv_inv:

  [| Group F; Group G; f ∈ gHom F G; x ∈ carrier F |] ==> fF x) = ρG f x

lemma ghomTr3:

  [| Group F; Group G; f ∈ gHom F G; x ∈ carrier F; y ∈ carrier F;
     f (Group.top F xF y)) = \<one>G |]
  ==> f x = f y

lemma iim_nonempty:

  [| Group F; Group G; f ∈ gHom F G; G » K  |] ==> iim F G f K  {}

lemma ghomTr4:

  [| Group F; Group G; f ∈ gHom F G; G » K  |] ==> F » iim F G f K 

lemma IdTr0:

  idmap (carrier G) ∈ gHom G G

lemma IdTr1:

  [| Group F; x ∈ carrier F |] ==> IF x = x

lemma IdTr2:

  Group F ==> gbijF,F IF

lemma Id_l_unit:

  [| Group G; gbijG,G f |] ==> IG oG f = f

8. gkernel

lemma gkernTr1:

  [| Group F; Group G; f ∈ gHom F G; x ∈ gkerF,G f |] ==> x ∈ carrier F

lemma gkernTr1_1:

  [| Group F; Group G; f ∈ gHom F G |] ==> gkerF,G f  carrier F

lemma gkernTr2:

  [| Group F; Group G; f ∈ gHom F G; x ∈ gkerF,G f; y ∈ gkerF,G f |]
  ==> Group.top F x y ∈ gkerF,G f

lemma gkernTr3:

  [| Group F; Group G; f ∈ gHom F G; x ∈ gkerF,G f |] ==> ρF x ∈ gkerF,G f

lemma gkernTr6:

  [| Group F; Group G; f ∈ gHom F G |] ==> \<one>F ∈ gkerF,G f

lemma gkernTr7:

  [| Group F; Group G; f ∈ gHom F G |] ==> F » gkerF,G f 

lemma gker_normal:

  [| Group F; Group G; f ∈ gHom F G |] ==> F \<triangleright> gkerF,G f

lemma Group_coim:

  [| Group F; Group G; f ∈ gHom F G |] ==> Group (F / gkerF,G f)

lemma gkern1:

  [| Group F; Ugp E; f ∈ gHom F E |] ==> gkerF,E f = carrier F

lemma gkern2:

  [| Group F; Group G; f ∈ gHom F G; ginjF,G f |] ==> gkerF,G f = {\<one>F}

lemma gkernTr9:

  [| Group F; Group G; f ∈ gHom F G; a ∈ carrier F; b ∈ carrier F |]
  ==> (gkerF,G fF a = gkerF,G fF b) = (f a = f b)

lemma gkernTr11:

  [| Group F; Group G; f ∈ gHom F G; a ∈ carrier F |]
  ==> iim F G f {f a} = gkerF,G fF a

lemma gbij_comp_bij:

  [| Group F; Group G; Group H; gbijF,G f; gbijG,H g |] ==> gbijF,H (g oF f)

lemma gbij_automorph:

  [| Group G; gbijG,G f; gbijG,G g |] ==> gbijG,G (g oG f)

lemma l_unit_gHom:

  [| Group F; Group G; f ∈ gHom F G |] ==> IG oF f = f

lemma r_unit_gHom:

  [| Group F; Group G; f ∈ gHom F G |] ==> f oF IF = f

9. Image

lemma inv_gHom:

  [| Group F; Group G; gbijF,G f |] ==> Ifn F G f ∈ gHom G F

lemma inv_gbijec_gbijec:

  [| Group F; Group G; gbijF,G f |] ==> gbijG,F (Ifn F G f)

lemma l_inv_gHom:

  [| Group F; Group G; gbijF,G f |] ==> Ifn F G f oF f = IF

lemma img_mult_closed:

  [| Group F; Group G; f ∈ gHom F G; uf ` carrier F; vf ` carrier F |]
  ==> Group.top G u vf ` carrier F

lemma img_unit_closed:

  [| Group F; Group G; f ∈ gHom F G |] ==> \<one>Gf ` carrier F

lemma imgTr7:

  [| Group F; Group G; f ∈ gHom F G; uf ` carrier F |] ==> ρG uf ` carrier F

lemma imgTr8:

  [| Group F; Group G; f ∈ gHom F G; F » H ; uf ` H; vf ` H |]
  ==> Group.top G u vf ` H

lemma imgTr9:

  [| Group F; Group G; f ∈ gHom F G; F » H ; uf ` H |] ==> ρG uf ` H

lemma imgTr10:

  [| Group F; Group G; f ∈ gHom F G; F » H  |] ==> \<one>Gf ` H

lemma imgTr11:

  [| Group F; Group G; f ∈ gHom F G; F » H  |] ==> G » f ` H 

lemma sg_gimg:

  [| Group F; Group G; f ∈ gHom F G |] ==> G » f ` carrier F 

lemma Group_Img:

  [| Group F; Group G; f ∈ gHom F G |] ==> Group (ImgF,G f)

lemma Img_carrier:

  [| Group F; Group G; f ∈ gHom F G |] ==> carrier (ImgF,G f) = f ` carrier F

lemma hom_to_Img:

  [| Group F; Group G; f ∈ gHom F G |] ==> f ∈ gHom F (ImgF,G f)

lemma gker_hom_to_img:

  [| Group F; Group G; f ∈ gHom F G |] ==> gkerF,ImgF,G f f = gkerF,G f

lemma Pj_im_subg:

  [| Group G; G » H ; G \<triangleright> K; K  H |]
  ==> Pj G K ` H = carrier (\<natural>GH / K)

lemma subg_Qsubg:

  [| G » H ; G \<triangleright> K; K  H |] ==> G / K » carrier (\<natural>H / K) 

10 incuded homomorphisms

lemma inducedhomTr:

  [| Group F; Group G; f ∈ gHom F G; S ∈ set_rcs F (gkerF,G f); s1.0S;
     s2.0S |]
  ==> f s1.0 = f s2.0

lemma induced_ghom_someTr:

  [| Group F; Group G; f ∈ gHom F G; X ∈ set_rcs F (gkerF,G f) |]
  ==> f (SOME xa. xaX) ∈ f ` carrier F

lemma induced_ghom_someTr1:

  [| Group F; Group G; f ∈ gHom F G; a ∈ carrier F |]
  ==> f (SOME xa. xa ∈ gkerF,G fF a) = f a

lemma inducedHOMTr0:

  [| Group F; Group G; f ∈ gHom F G; a ∈ carrier F |]
  ==> (f¨F,G) (gkerF,G fF a) = f a

lemma inducedHOMTr0_1:

  [| Group F; Group G; f ∈ gHom F G |]
  ==> f¨F,G ∈ set_rcs F (gkerF,G f) -> carrier G

lemma inducedHOMTr0_2:

  [| Group F; Group G; f ∈ gHom F G |]
  ==> f¨F,G ∈ set_rcs F (gkerF,G f) -> f ` carrier F

lemma inducedHom:

  [| Group F; Group G; f ∈ gHom F G |] ==> f¨F,G ∈ gHom (F / gkerF,G f) G

lemma induced_ghom_ginjec:

  [| Group F; Group G; f ∈ gHom F G |] ==> ginj(F / gkerF,G f),G (f¨F,G)

lemma inducedhomgsurjec:

  [| Group F; Group G; gsurjF,G f |] ==> gsurj(F / gkerF,G f),G (f¨F,G)

lemma homomtr:

  [| Group F; Group G; f ∈ gHom F G |] ==> f¨F,G ∈ gHom (F / gkerF,G f) (ImgF,G f)

lemma homom2img:

  [| Group F; Group G; f ∈ gHom F G |]
  ==> f¨F,ImgF,G f ∈ gHom (F / gkerF,G f) (ImgF,G f)

lemma homom2img1:

  [| Group F; Group G; f ∈ gHom F G; X ∈ set_rcs F (gkerF,G f) |]
  ==> (f¨F,ImgF,G f) X = (f¨F,G) X

11 Homomorphism therems

lemma iotahomTr0:

  [| Group G; G » H ; hH |] ==> (ι(\<natural>GH)) h = h

lemma iotahom:

  [| Group G; G » H ; G \<triangleright> N |]
  ==> ι(\<natural>GH) ∈ gHom (\<natural>GH) (\<natural>GH \<struct>G N)

lemma iotaTr0:

  [| Group G; G » H ; G \<triangleright> N |]
  ==> ginj(\<natural>GH),(\<natural>GH \<struct>G N) ι(\<natural>GH)

theorem homomthm1:

  [| Group F; Group G; f ∈ gHom F G |]
  ==> gbij(F / gkerF,G f),ImgF,G f (f¨F,ImgF,G f)

lemma isomTr0:

  Group F ==> FF

lemma isomTr1:

  [| Group F; Group G; FG |] ==> GF

lemma isomTr2:

  [| Group F; Group G; Group H; FG; GH |] ==> FH

lemma gisom1:

  [| Group F; Group G; f ∈ gHom F G |] ==> (F / gkerF,G f) ≅ (ImgF,G f)

lemma homomth2Tr0:

  [| Group F; Group G; f ∈ gHom F G; G \<triangleright> N |]
  ==> F \<triangleright> iim F G f N

lemma kern_comp_gHom:

  [| Group F; Group G; gsurjF,G f; G \<triangleright> N |]
  ==> gkerF,(G / N) (Pj G N oF f) = iim F G f N

lemma QgrpUnit_1:

  [| Group G; Ugp E; G \<triangleright> H; (G / H) ≅ E |] ==> carrier G = H

lemma QgrpUnit_2:

  [| Group G; Ugp E; G \<triangleright> H; carrier G = H |] ==> (G / H) ≅ E

lemma QgrpUnit_3:

  [| Group G; Ugp E; G » H ; G » H1.0 ; \<natural>GH \<triangleright> H1.0;
     (\<natural>GH / H1.0) ≅ E |]
  ==> H = H1.0

lemma QgrpUnit_4:

  [| Group G; Ugp E; G » H ; G » H1.0 ; \<natural>GH \<triangleright> H1.0;
     ¬ (\<natural>GH / H1.0) ≅ E |]
  ==> H  H1.0

lemma QmpTr0:

  [| G » H ; G » N ; H  N; a ∈ carrier G |] ==> (QmG H,N) (Ha) = Na

lemma QmpTr1:

  [| G » H ; G » N ; H  N; a ∈ carrier G; b ∈ carrier G; Ha = Hb |]
  ==> Na = Nb

lemma QmpTr2:

  [| G » H ; G » N ; H  N; X ∈ carrier (G / H) |]
  ==> (QmG H,N) X ∈ carrier (G / N)

lemma QmpTr2_1:

  [| G » H ; G » N ; H  N |] ==> QmG H,N ∈ carrier (G / H) -> carrier (G / N)

lemma QmpTr3:

  [| G \<triangleright> H; G \<triangleright> N; H  N; X ∈ carrier (G / H);
     Y ∈ carrier (G / H) |]
  ==> (QmG H,N) (c_top G H X Y) = c_top G N ((QmG H,N) X) ((QmG H,N) Y)

lemma Gp_s_mult_nsg:

  [| G \<triangleright> H; G \<triangleright> N; H  N; aN |]
  ==> H\<natural>N a = Ha

lemma QmpTr5:

  [| G \<triangleright> H; G \<triangleright> N; H  N; X ∈ carrier (G / H);
     Y ∈ carrier (G / H) |]
  ==> (QmG H,N) (Group.top (G / H) X Y) =
      Group.top (G / N) ((QmG H,N) X) ((QmG H,N) Y)

lemma QmpTr:

  [| G \<triangleright> H; G \<triangleright> N; H  N |]
  ==> QmG H,N ∈ gHom (G / H) (G / N)

lemma Qmpgsurjec:

  [| G \<triangleright> H; G \<triangleright> N; H  N |]
  ==> gsurj(G / H),(G / N) (QmG H,N)

lemma gkerQmp:

  [| G \<triangleright> H; G \<triangleright> N; H  N |]
  ==> gker(G / H),(G / N) (QmG H,N) = carrier (\<natural>N / H)

theorem homom2:

  [| G \<triangleright> H; G \<triangleright> N; H  N |]
  ==> gbij(G / H / carrier (\<natural>N / H)),(G / N) (QmG H,N¨(G / H),(G / N))

12. Isomorphims

theorem isom2:

  [| G \<triangleright> H; G \<triangleright> N; H  N |]
  ==> (G / H / carrier (\<natural>N / H)) ≅ (G / N)

theorem homom3:

  [| Group F; Group G; G \<triangleright> N; gsurjF,G f; N1.0 = iim F G f N |]
  ==> (F / N1.0) ≅ (G / N)

lemma homom3Tr1:

  [| G » H ; G \<triangleright> N |]
  ==> HN =
      gker(\<natural>H),(\<natural>H \<struct>G N /
                         N) (Pj (\<natural>H \<struct>G N)
                              N o(\<natural>H) ι(\<natural>H))

An automorphism groups

lemma automgroupTr1:

  [| Group G; gbijG,G f; gbijG,G g; gbijG,G h |] ==> h oG g oG f = h oG (g oG f)

lemma automgroup:

  Group G ==> Group (automg G)

complete system of representatives

lemma gcsrp_func:

  G » H  ==> gcsrp_map G H ∈ set_rcs G H -> UNIV

lemma gcsrp_func1:

  G » H  ==> gcsrp_map G H ∈ set_rcs G H -> gcsrp_map G H ` set_rcs G H

lemma gcsrp_map_bij:

  G » H  ==> bij_to (gcsrp_map G H) (set_rcs G H) (gcsrp_map G H ` set_rcs G H)

lemma image_gcsrp:

  G » H  ==> gcsrp G H (gcsrp_map G H ` set_rcs G H)

lemma gcsrp_exists:

  G » H  ==> ∃S. gcsrp G H S

lemma gcsrp_top_closed:

  [| Group G; G \<triangleright> N; a ∈ gcsrp_map G N ` set_rcs G N;
     b ∈ gcsrp_map G N ` set_rcs G N |]
  ==> gcsrp_top G N a b ∈ gcsrp_map G N ` set_rcs G N

lemma gcsrp_tassoc:

  [| Group G; G \<triangleright> N; a ∈ gcsrp_map G N ` set_rcs G N;
     b ∈ gcsrp_map G N ` set_rcs G N; c ∈ gcsrp_map G N ` set_rcs G N |]
  ==> gcsrp_top G N (gcsrp_top G N a b) c = gcsrp_top G N a (gcsrp_top G N b c)

lemma gcsrp_l_one:

  [| Group G; G \<triangleright> N; a ∈ gcsrp_map G N ` set_rcs G N |]
  ==> gcsrp_top G N (gcsrp_one G N) a = a

lemma gcsrp_l_i:

  [| G \<triangleright> N; a ∈ gcsrp_map G N ` set_rcs G N |]
  ==> gcsrp_top G N (gcsrp_iop G N a) a = gcsrp_one G N

lemma gcsrp_i_closed:

  [| G \<triangleright> N; a ∈ gcsrp_map G N ` set_rcs G N |]
  ==> gcsrp_iop G N a ∈ gcsrp_map G N ` set_rcs G N

lemma Group_Gcsrp:

  G \<triangleright> N ==> Group (Gcsrp G N)

lemma gcsrp_map_gbijec:

  G \<triangleright> N ==> gbij(G / N),Gcsrp G N gcsrp_map G N

lemma Qg_equiv_Gcsrp:

  G \<triangleright> N ==> (G / N) ≅ Gcsrp G N

14. Zassenhaus

lemma homom4Tr1:

  [| G \<triangleright> N; G » H  |] ==> Group (\<natural>H \<struct>G N / N)

lemma homom3Tr2:

  [| Group G; G » H ; G \<triangleright> N |]
  ==> gsurj(\<natural>GH),(\<natural>GH \<struct>G N /
                           N) (Pj (\<natural>GH \<struct>G N)
                                N o(\<natural>GH) ι(\<natural>GH))

theorem homom4:

  [| Group G; G \<triangleright> N; G » H  |]
  ==> gbij(\<natural>GH /
           (HN)),(\<natural>GH \<struct>G N /
                     N) (Pj (\<natural>GH \<struct>G N)
                          N o(\<natural>GH) ι(\<natural>GH)¨(\<natural>GH),(\<natural>GH \<struct>G N /
                                    N))

lemma homom4_2:

  [| G \<triangleright> N; G » H  |] ==> Group (\<natural>H / (HN))

lemma isom4:

  [| Group G; G \<triangleright> N; G » H  |]
  ==> (\<natural>GH / (HN)) ≅ (\<natural>GN \<struct>G H / N)

lemma ZassenhausTr5:

  [| Group G; G » H ; G » H1.0 ; G » K ; G » K1.0 ;
     \<natural>GH \<triangleright> H1.0; \<natural>GK \<triangleright> K1.0 |]
  ==> (\<natural>GHK / (H1.0K \<struct>G HK1.0)) ≅
      (\<natural>GH1.0 \<struct>G HK / (H1.0 \<struct>G HK1.0))

lemma ZassenhausTr5_1:

  [| Group G; G » H ; G » H1.0 ; G » K ; G » K1.0 ;
     \<natural>GH \<triangleright> H1.0; \<natural>GK \<triangleright> K1.0 |]
  ==> (\<natural>GKH / (K1.0H \<struct>G KH1.0)) ≅
      (\<natural>GK1.0 \<struct>G KH / (K1.0 \<struct>G KH1.0))

lemma ZassenhausTr5_2:

  [| Group G; G » H ; G » H1.0 ; G » K ; G » K1.0 ;
     \<natural>GH \<triangleright> H1.0; \<natural>GK \<triangleright> K1.0 |]
  ==> \<natural>GHK / (H1.0K \<struct>G HK1.0) =
      \<natural>GKH / (K1.0H \<struct>G KH1.0)

lemma ZassenhausTr6_1:

  [| Group G; G » H ; G » H1.0 ; G » K ; G » K1.0 ;
     \<natural>GH \<triangleright> H1.0; \<natural>GK \<triangleright> K1.0 |]
  ==> Group (\<natural>GHK / (H1.0K \<struct>G HK1.0))

lemma ZassenhausTr6_2:

  [| Group G; G » H ; G » H1.0 ; G » K ; G » K1.0 ;
     \<natural>GH \<triangleright> H1.0; \<natural>GK \<triangleright> K1.0 |]
  ==> Group (\<natural>GH1.0 \<struct>G HK / (H1.0 \<struct>G HK1.0))

lemma ZassenhausTr6_3:

  [| Group G; G » H ; G » H1.0 ; G » K ; G » K1.0 ;
     \<natural>GH \<triangleright> H1.0; \<natural>GK \<triangleright> K1.0 |]
  ==> Group (\<natural>GK1.0 \<struct>G KH / (K1.0 \<struct>G KH1.0))

theorem Zassenhaus:

  [| Group G; G » H ; G » H1.0 ; G » K ; G » K1.0 ;
     \<natural>GH \<triangleright> H1.0; \<natural>GK \<triangleright> K1.0 |]
  ==> (\<natural>GH1.0 \<struct>G HK / (H1.0 \<struct>G HK1.0)) ≅
      (\<natural>GK1.0 \<struct>G KH / (K1.0 \<struct>G KH1.0))

15. chain of groups I

lemma d_gchainTr0:

  [| 0 < n; d_gchain G n f; k  n - 1 |] ==> f (Suc k)  f k

lemma d_gchain_mem_sg:

  d_gchain G n f ==> ∀in. G » f i 

lemma d_gchain_pre:

  d_gchain G (Suc n) f ==> d_gchain G n f

lemma d_gchainTr1:

  0 < n --> (∀f. d_gchain G n f --> (∀ln. ∀jn. l < j --> f j  f l))

lemma d_gchainTr2:

  [| 0 < n; d_gchain G n f; l  n; j  n; l  j |] ==> f j  f l

lemma im_d_gchainTr1:

  [| d_gchain G n f; f lf ` {i. i  n} - {f 0} |]
  ==> f (LEAST j. f jf ` {i. i  n} - {f 0}) ∈ f ` {i. i  n} - {f 0}

lemma im_d_gchainTr1_0:

  [| d_gchain G n f; f lf ` {i. i  n} - {f 0} |]
  ==> 0 < (LEAST j. f jf ` {i. i  n} - {f 0})

lemma im_d_gchainTr1_1:

  [| d_gchain G n f; ∃i. f if ` {i. i  n} - {f 0} |]
  ==> f (LEAST j. f jf ` {i. i  n} - {f 0}) ∈ f ` {i. i  n} - {f 0}

lemma im_d_gchainsTr1_2:

  [| d_gchain G n f; i  n; f if ` {i. i  n} - {f 0} |]
  ==> (LEAST j. f jf ` {i. i  n} - {f 0})  i

lemma im_d_gchainsTr1_3:

  [| d_gchain G n f; ∃in. f if ` {i. i  n} - {f 0};
     k < (LEAST j. f jf ` {i. i  n} - {f 0}) |]
  ==> f k = f 0

lemma im_gdchainsTr1_4:

  [| d_gchain G n f; ∃vf ` {i. i  n}. v  {f 0};
     i < (LEAST j. f jf ` {i. i  n} ∧ f j  f 0) |]
  ==> f i = f 0

lemma im_d_gchainsTr1_5:

  [| 0 < n; d_gchain G n f; i  n; f if ` {i. i  n} - {f 0};
     (LEAST j. f jf ` {i. i  n} - {f 0}) = j |]
  ==> f ` {i. i  j - Suc 0} = {f 0}

lemma im_d_gchains1:

  [| 0 < n; d_gchain G n f; i  n; f if ` {i. i  n} - {f 0};
     (LEAST j. f jf ` {i. i  n} - {f 0}) = j |]
  ==> f ` {i. i  n} = {f 0} ∪ {f i |i. j  ii  n}

lemma im_d_gchains1_1:

  [| d_gchain G n f; f n  f 0 |]
  ==> f ` {i. i  n} =
      {f 0} ∪ {f i |i. (LEAST j. f jf ` {i. i  n} - {f 0})  ii  n}

lemma d_gchains_leastTr:

  [| d_gchain G n f; f n  f 0 |]
  ==> (LEAST j. f jf ` {i. i  n} - {f 0}) ∈ {i. i  n} ∧
      f (LEAST j. f jf ` {i. i  n} - {f 0})  f 0

lemma im_d_gchainTr2:

  [| d_gchain G n f; j  n; f j  f 0 |] ==> ∀in. f 0 = f i --> ¬ j  i

lemma D_gchain_pre:

  D_gchain G (Suc n) f ==> D_gchain G n f

lemma D_gchain0:

  [| D_gchain G n f; i  n; j  n; i < j |] ==> f j  f i

lemma D_gchain1:

  D_gchain G n f ==> inj_on f {i. i  n}

lemma card_im_D_gchain:

  [| 0 < n; D_gchain G n f |] ==> card (f ` {i. i  n}) = Suc n

lemma w_cmpser_gr:

  [| 0 < r; w_cmpser G r f; i  r |] ==> G » f i 

lemma w_cmpser_ns:

  [| 0 < r; w_cmpser G r f; i  r - 1 |]
  ==> \<natural>f i \<triangleright> f (Suc i)

lemma w_cmpser_pre:

  w_cmpser G (Suc n) f ==> w_cmpser G n f

lemma W_cmpser_pre:

  W_cmpser G (Suc n) f ==> W_cmpser G n f

lemma td_gchain_n:

  [| td_gchain G n f; carrier G  {\<one>} |] ==> 0 < n

16. Existence of reduced chain

lemma D_gchain_is_d_gchain:

  D_gchain G n f ==> d_gchain G n f

lemma joint_d_gchains:

  [| d_gchain G n f; d_gchain G m g; g 0  f n |]
  ==> d_gchain G (Suc (n + m)) (jointfun n f m g)

lemma joint_D_gchains:

  [| D_gchain G n f; D_gchain G m g; g 0  f n |]
  ==> D_gchain G (Suc (n + m)) (jointfun n f m g)

lemma w_cmpser_is_d_gchain:

  w_cmpser G n f ==> d_gchain G n f

lemma joint_w_cmpser:

  [| w_cmpser G n f; w_cmpser G m g; \<natural>f n \<triangleright> g 0 |]
  ==> w_cmpser G (Suc (n + m)) (jointfun n f m g)

lemma W_cmpser_is_D_gchain:

  W_cmpser G n f ==> D_gchain G n f

lemma W_cmpser_is_w_cmpser:

  W_cmpser G n f ==> w_cmpser G n f

lemma tw_cmpser_is_w_cmpser:

  tw_cmpser G n f ==> w_cmpser G n f

lemma tW_cmpser_is_W_cmpser:

  tW_cmpser G n f ==> W_cmpser G n f

lemma joint_W_cmpser:

  [| W_cmpser G n f; W_cmpser G m g; \<natural>f n \<triangleright> g 0;
     g 0  f n |]
  ==> W_cmpser G (Suc (n + m)) (jointfun n f m g)

lemma joint_d_gchain_n0:

  [| d_gchain G n f; d_gchain G 0 g; g 0  f n |]
  ==> d_gchain G (Suc n) (jointfun n f 0 g)

lemma joint_D_gchain_n0:

  [| D_gchain G n f; D_gchain G 0 g; g 0  f n |]
  ==> D_gchain G (Suc n) (jointfun n f 0 g)

lemma joint_w_cmpser_n0:

  [| w_cmpser G n f; w_cmpser G 0 g; \<natural>f n \<triangleright> g 0 |]
  ==> w_cmpser G (Suc n) (jointfun n f 0 g)

lemma joint_W_cmpser_n0:

  [| W_cmpser G n f; W_cmpser G 0 g; \<natural>f n \<triangleright> g 0;
     g 0  f n |]
  ==> W_cmpser G (Suc n) (jointfun n f 0 g)

lemma compseriesTr0:

  [| compseries G n f; i  n |] ==> G » f i 

lemma compseriesTr1:

  compseries G n f ==> tW_cmpser G n f

lemma compseriesTr2:

  compseries G n f ==> f 0 = carrier G

lemma compseriesTr3:

  compseries G n f ==> f n = {\<one>}

lemma compseriesTr4:

  compseries G n f ==> w_cmpser G n f

lemma im_jointfun1Tr1:

  ln. G » f l  ==> f ∈ {i. i  n} -> Collect (sg G)

lemma Nset_Suc_im:

  l≤Suc n. G » f l  ==> insert (f (Suc n)) (f ` {i. i  n}) = f ` {i. i  Suc n}

lemma LeastTr0:

  i < (LEAST l. P l) ==> ¬ P i

lemma funeq_LeastTr1:

  [| ∀ln. G » f l ; ∀ln. G » g l ; l < (LEAST k. NfuncPair_neq_at f g k) |]
  ==> f l = g l

lemma funeq_LeastTr1_1:

  [| ∀ln. G » f l ; ∀ln. G » g l ; l < (LEAST k. f k  g k) |] ==> f l = g l

lemma Nfunc_LeastTr2_1:

  [| i  n; ∀ln. G » f l ; ∀ln. G » g l ; NfuncPair_neq_at f g i |]
  ==> NfuncPair_neq_at f g (LEAST k. NfuncPair_neq_at f g k)

lemma Nfunc_LeastTr2_2:

  [| i  n; ∀ln. G » f l ; ∀ln. G » g l ; NfuncPair_neq_at f g i |]
  ==> (LEAST k. NfuncPair_neq_at f g k)  i

lemma Nfunc_LeastTr2_2_1:

  [| i  n; ∀ln. G » f l ; ∀ln. G » g l ; f i  g i |]
  ==> (LEAST k. f k  g k)  i

lemma Nfunc_LeastTr2_3:

  [| ∀ln. G » f l ; ∀ln. G » g l ; i  n; f i  g i |]
  ==> f (LEAST k. f k  g k)  g (LEAST k. f k  g k)

lemma Nfunc_LeastTr2_4:

  [| ∀ln. G » f l ; ∀ln. G » g l ; i  n; f i  g i |]
  ==> (LEAST k. f k  g k)  n

lemma Nfunc_LeastTr2_5:

  [| ∀ln. G » f l ; ∀ln. G » g l ; ∃in. f i  g i |]
  ==> f (LEAST k. f k  g k)  g (LEAST k. f k  g k)

lemma Nfunc_LeastTr2_6:

  [| ∀ln. G » f l ; ∀ln. G » g l ; ∃in. f i  g i |]
  ==> (LEAST k. f k  g k)  n

lemma Nfunc_Least_sym:

  [| ∀ln. G » f l ; ∀ln. G » g l ; ∃in. f i  g i |]
  ==> (LEAST k. f k  g k) = (LEAST k. g k  f k)

lemma Nfunc_iNJTr:

  [| inj_on g {i. i  n}; i  n; j  n; i < j |] ==> g i  g j

lemma Nfunc_LeastTr2_7:

  [| ∀ln. G » f l ; ∀ln. G » g l ; inj_on g {i. i  n}; ∃in. f i  g i;
     f k = g (LEAST k. f k  g k) |]
  ==> (LEAST k. f k  g k) < k

lemma Nfunc_LeastTr2_8:

  [| ∀ln. G » f l ; ∀ln. G » g l ; inj_on g {i. i  n}; ∃in. f i  g i;
     f ` {i. i  n} = g ` {i. i  n} |]
  ==> ∃k∈nset (Suc (LEAST i. f i  g i)) n. f k = g (LEAST i. f i  g i)

lemma ex_redchainTr1:

  [| d_gchain G n f; D_gchain G (card (f ` {i. i  n}) - Suc 0) g;
     g ` {i. i  card (f ` {i. i  n}) - Suc 0} = f ` {i. i  n} |]
  ==> g (card (f ` {i. i  n}) - Suc 0) = f n

lemma ex_redchainTr1_1:

  [| d_gchain G n f; D_gchain G (card (f ` {i. i  n}) - Suc 0) g;
     g ` {i. i  card (f ` {i. i  n}) - Suc 0} = f ` {i. i  n} |]
  ==> g 0 = f 0

lemma ex_redchainTr2:

  d_gchain G (Suc n) f ==> D_gchain G 0 (constmap {0} {f (Suc n)})

lemma last_mem_excluded:

  [| d_gchain G (Suc n) f; f n  f (Suc n) |] ==> f (Suc n)  f ` {i. i  n}

lemma ex_redchainTr4:

  [| d_gchain G (Suc n) f; f n  f (Suc n) |]
  ==> card (f ` {i. i  Suc n}) = Suc (card (f ` {i. i  n}))

lemma ex_redchainTr5:

  d_gchain G n f ==> 0 < card (f ` {i. i  n})

lemma ex_redchainTr6:

  f. d_gchain G n f -->
      (∃g. D_gchain G (card (f ` {i. i  n}) - 1) gg ` {i. i  card (f ` {i. i  n}) - 1} = f ` {i. i  n})

lemma ex_redchain:

  d_gchain G n f
  ==> ∃g. D_gchain G (card (f ` {i. i  n}) - 1) gg ` {i. i  card (f ` {i. i  n}) - 1} = f ` {i. i  n}

lemma const_W_cmpser:

  d_gchain G (Suc n) f ==> W_cmpser G 0 (constmap {0} {f (Suc n)})

lemma ex_W_cmpserTr0m:

  f. w_cmpser G m f -->
      (∃g. W_cmpser G (card (f ` {i. i  m}) - 1) gg ` {i. i  card (f ` {i. i  m}) - 1} = f ` {i. i  m})

lemma ex_W_cmpser:

  w_cmpser G m f
  ==> ∃g. W_cmpser G (card (f ` {i. i  m}) - 1) gg ` {i. i  card (f ` {i. i  m}) - 1} = f ` {i. i  m}

17. Existence of reduced chain and composition series

lemma ex_W_cmpserTr3m1:

  [| tw_cmpser G m f; W_cmpser G (card (f ` {i. i  m}) - 1) g;
     g ` {i. i  card (f ` {i. i  m}) - 1} = f ` {i. i  m} |]
  ==> tW_cmpser G (card (f ` {i. i  m}) - 1) g

lemma ex_W_cmpserTr3m:

  tw_cmpser G m f
  ==> ∃g. tW_cmpser G (card (f ` {i. i  m}) - 1) gg ` {i. i  card (f ` {i. i  m}) - 1} = f ` {i. i  m}

lemma red_chainTr0m1_1:

  tw_cmpser G m f
  ==> (SOME g. g ∈ {h. red_ch_cd G f m h}) ∈ {h. red_ch_cd G f m h}

lemma red_chain_m:

  tw_cmpser G m f
  ==> tW_cmpser G (card (f ` {i. i  m}) - 1) (red_chain G m f) ∧
      red_chain G m f ` {i. i  card (f ` {i. i  m}) - 1} = f ` {i. i  m}

18. Chain of groups II

lemma Gchain_pre:

  Gchain (Suc n) g ==> Gchain n g

lemma isom_unit:

  [| G » H ; G » K ; H = {\<one>} |]
  ==> (\<natural>H) ≅ (\<natural>K) --> K = {\<one>}

lemma isom_gch_unitsTr4:

  [| Group F; Group G; Ugp E; FG; FE |] ==> GE

lemma isom_gch_cmp:

  [| Gchain n g; Gchain n h; f1.0 ∈ {i. i  n} -> {i. i  n};
     f2.0 ∈ {i. i  n} -> {i. i  n}; isom_Gchains n (cmp f2.0 f1.0) g h |]
  ==> isom_Gchains n f1.0 g (cmp h f2.0)

lemma isom_gch_transp:

  [| Gchain n f; i  n; j  n; i < j |]
  ==> isom_Gchains n (transpos i j) f (cmp f (transpos i j))

lemma isom_gch_units_transpTr0:

  [| Ugp E; Gchain n g; Gchain n h; i  n; j  n; i < j;
     isom_Gchains n (transpos i j) g h |]
  ==> {i. i  ng iE} - {i, j} = {i. i  nh iE} - {i, j}

lemma isom_gch_units_transpTr1:

  [| Ugp E; Gchain n g; i  n; j  n; g jE; i  j |]
  ==> insert j ({i. i  ng iE} - {i, j}) = {i. i  ng iE} - {i}

lemma isom_gch_units_transpTr2:

  [| Ugp E; Gchain n g; i  n; j  n; i < j; g iE |]
  ==> {i. i  ng iE} = insert i ({i. i  ng iE} - {i})

lemma isom_gch_units_transpTr3:

  [| Ugp E; Gchain n g; i  n |] ==> finite ({i. i  ng iE} - {i})

lemma isom_gch_units_transpTr4:

  [| Ugp E; Gchain n g; i  n |] ==> finite ({i. i  ng iE} - {i, j})

lemma isom_gch_units_transpTr5_1:

  [| Ugp E; Gchain n g; Gchain n h; i  n; j  n; i < j;
     isom_Gchains n (transpos i j) g h |]
  ==> g ih j

lemma isom_gch_units_transpTr5_2:

  [| Ugp E; Gchain n g; Gchain n h; i  n; j  n; i < j;
     isom_Gchains n (transpos i j) g h |]
  ==> g jh i

lemma isom_gch_units_transpTr6:

  [| Gchain n g; i  n |] ==> Group (g i)

lemma isom_gch_units_transpTr7:

  [| Ugp E; i  n; j  n; g jh i; Group (h i); Group (g j); ¬ g jE |]
  ==> ¬ h iE

lemma isom_gch_units_transpTr8_1:

  [| Ugp E; Gchain n g; i  n; j  n; g iE; ¬ g jE |]
  ==> {i. i  ng iE} = {i. i  ng iE} - {j}

lemma isom_gch_units_transpTr8_2:

  [| Ugp E; Gchain n g; i  n; j  n; ¬ g iE; ¬ g jE |]
  ==> {i. i  ng iE} = {i. i  ng iE} - {i, j}

lemma isom_gch_units_transp:

  [| Ugp E; Gchain n g; Gchain n h; i  n; j  n; i < j;
     isom_Gchains n (transpos i j) g h |]
  ==> card {i. i  ng iE} = card {i. i  nh iE}

lemma TR_isom_gch_units:

  [| Ugp E; Gchain n f; i  n; j  n; i < j |]
  ==> card {k. k  nf kE} = card {k. k  n ∧ cmp f (transpos i j) kE}

lemma TR_isom_gch_units_1:

  [| Ugp E; Gchain n f; i  n; j  n; i < j |]
  ==> card {k. k  nf kE} = card {k. k  nf (transpos i j k) ≅ E}

lemma isom_tgch_unitsTr0_1:

  [| Ugp E; Gchain (Suc n) g; g (Suc n) ≅ E |]
  ==> {i. i  Suc ng iE} = insert (Suc n) {i. i  ng iE}

lemma isom_tgch_unitsTr0_2:

  Ugp E ==> finite {i. i  ng iE}

lemma isom_tgch_unitsTr0_3:

  [| Ugp E; Gchain (Suc n) g; ¬ g (Suc n) ≅ E |]
  ==> {i. i  Suc ng iE} = {i. i  ng iE}

lemma isom_tgch_unitsTr0:

  [| Ugp E; card {i. i  ng iE} = card {i. i  nh iE};
     Gchain (Suc n) g ∧ Gchain (Suc n) h ∧ Gch_bridge (Suc n) g h f;
     f (Suc n) = Suc n |]
  ==> card {i. i  Suc ng iE} = card {i. i  Suc nh iE}

lemma isom_gch_unitsTr1_1:

  [| Ugp E; Gchain (Suc n) g ∧ Gchain (Suc n) h ∧ Gch_bridge (Suc n) g h f;
     f (Suc n) = Suc n |]
  ==> Gchain n g ∧ Gchain n h ∧ Gch_bridge n g h f

lemma isom_gch_unitsTr1_2:

  [| Ugp E; f (Suc n)  Suc n; inj_on f {i. i  Suc n}; ∀l≤Suc n. f l  Suc n |]
  ==> cmp (transpos (f (Suc n)) (Suc n)) f (Suc n) = Suc n

lemma isom_gch_unitsTr1_3:

  [| Ugp E; f (Suc n)  Suc n; ∀l≤Suc n. f l  Suc n; inj_on f {i. i  Suc n} |]
  ==> inj_on (cmp (transpos (f (Suc n)) (Suc n)) f) {i. i  Suc n}

lemma isom_gch_unitsTr1_4:

  [| Ugp E; f (Suc n)  Suc n; inj_on f {i. i  Suc n}; ∀l≤Suc n. f l  Suc n |]
  ==> inj_on (cmp (transpos (f (Suc n)) (Suc n)) f) {i. i  n}

lemma isom_gch_unitsTr1_5:

  [| Ugp E; Gchain (Suc n) g ∧ Gchain (Suc n) h ∧ Gch_bridge (Suc n) g h f;
     f (Suc n)  Suc n |]
  ==> Gchain n g ∧
      Gchain n (cmp h (transpos (f (Suc n)) (Suc n))) ∧
      Gch_bridge n g (cmp h (transpos (f (Suc n)) (Suc n)))
       (cmp (transpos (f (Suc n)) (Suc n)) f)

lemma isom_gch_unitsTr1_6:

  [| Ugp E; f (Suc n)  Suc n;
     Gchain (Suc n) g ∧ Gchain (Suc n) h ∧ Gch_bridge (Suc n) g h f |]
  ==> Gchain (Suc n) g ∧
      Gchain (Suc n) (cmp h (transpos (f (Suc n)) (Suc n))) ∧
      Gch_bridge (Suc n) g (cmp h (transpos (f (Suc n)) (Suc n)))
       (cmp (transpos (f (Suc n)) (Suc n)) f)

lemma isom_gch_unitsTr1_7_0:

  [| Gchain (Suc n) h; k  Suc n; k  Suc n |]
  ==> Gchain (Suc n) (cmp h (transpos k (Suc n)))

lemma isom_gch_unitsTr1_7_1:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n |]
  ==> {i. i  Suc n ∧ cmp h (transpos k (Suc n)) iE} - {k, Suc n} =
      {i. i  Suc nh iE} - {k, Suc n}

lemma isom_gch_unitsTr1_7_2:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n; h (Suc n) ≅ E |]
  ==> cmp h (transpos k (Suc n)) kE

lemma isom_gch_unitsTr1_7_3:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n; h kE |]
  ==> cmp h (transpos k (Suc n)) (Suc n) ≅ E

lemma isom_gch_unitsTr1_7_4:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n; ¬ h (Suc n) ≅ E |]
  ==> ¬ cmp h (transpos k (Suc n)) kE

lemma isom_gch_unitsTr1_7_5:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n; ¬ h kE |]
  ==> ¬ cmp h (transpos k (Suc n)) (Suc n) ≅ E

lemma isom_gch_unitsTr1_7_6:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n; h (Suc n) ≅ E; h kE |]
  ==> {i. i  Suc nh iE} =
      insert k (insert (Suc n) ({i. i  Suc nh iE} - {k, Suc n}))

lemma isom_gch_unitsTr1_7_7:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n; h (Suc n) ≅ E; ¬ h kE |]
  ==> {i. i  Suc nh iE} =
      insert (Suc n) ({i. i  Suc nh iE} - {k, Suc n})

lemma isom_gch_unitsTr1_7_8:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n; ¬ h (Suc n) ≅ E; h kE |]
  ==> {i. i  Suc nh iE} = insert k ({i. i  Suc nh iE} - {k, Suc n})

lemma isom_gch_unitsTr1_7_9:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n; ¬ h (Suc n) ≅ E; ¬ h kE |]
  ==> {i. i  Suc nh iE} = {i. i  Suc nh iE} - {k, Suc n}

lemma isom_gch_unitsTr1_7:

  [| Ugp E; Gchain (Suc n) h; k  Suc n; k  Suc n |]
  ==> card {i. i  Suc n ∧ cmp h (transpos k (Suc n)) iE} =
      card {i. i  Suc nh iE}

lemma isom_gch_unitsTr1:

  Ugp E
  ==> ∀g h f.
         Gchain n g ∧ Gchain n h ∧ Gch_bridge n g h f -->
         card {i. i  ng iE} = card {i. i  nh iE}

lemma isom_gch_units:

  [| Ugp E; Gchain n g; Gchain n h; Gch_bridge n g h f |]
  ==> card {i. i  ng iE} = card {i. i  nh iE}

lemma isom_gch_units_1:

  [| Ugp E; Gchain n g; Gchain n h; ∃f. Gch_bridge n g h f |]
  ==> card {i. i  ng iE} = card {i. i  nh iE}

19. Jordan Hoelder theorem

rfn_tools. tools to treat refinement of a cmpser, rtos

lemma rfn_tool1:

  [| 0 < r; k = i * r + j; j < r |] ==> k div r = i

lemma pos_mult_pos:

  [| 0 < r; 0 < s |] ==> 0 < r * s

lemma rfn_tool1_1:

  [| 0 < r; j < r |] ==> (i * r + j) div r = i

lemma rfn_tool2:

  a < s ==> a  s - Suc 0

lemma rfn_tool3:

  0  m ==> m + n - n = m

lemma rfn_tool11:

  [| 0 < b; a  b - Suc 0 |] ==> a < b

lemma rfn_tool12:

  [| 0 < s; i mod s = s - 1 |] ==> Suc (i div s) = Suc i div s

lemma rfn_tool12_1:

  [| 0 < s; l mod s < s - 1 |] ==> Suc (l mod s) = Suc l mod s

lemma rfn_tool12_2:

  [| 0 < s; i mod s = s - Suc 0 |] ==> Suc i mod s = 0

lemma rfn_tool13:

  [| 0 < r; a = b |] ==> a mod r = b mod r

lemma rfn_tool13_1:

  [| 0 < r; a = b |] ==> a div r = b div r

lemma div_Tr1:

  [| 0 < r; 0 < s; l  s * r |] ==> l div s  r

lemma div_Tr2:

  [| 0 < r; 0 < s; l < s * r |] ==> l div s  r - Suc 0

lemma div_Tr3:

  [| 0 < r; 0 < s; l < s * r |] ==> Suc (l div s)  r

lemma div_Tr3_1:

  [| 0 < r; 0 < s; l mod s = s - 1 |] ==> Suc l div s = Suc (l div s)

lemma div_Tr3_2:

  [| 0 < r; 0 < s; l mod s < s - 1 |] ==> l div s = Suc l div s

lemma mod_div_injTr:

  [| 0 < r; x mod r = y mod r; x div r = y div r |] ==> x = y

lemma rtos_hom0:

  [| 0 < r; 0 < s; i  r * s - Suc 0 |] ==> i div s < r

lemma rtos_hom1:

  [| 0 < r; 0 < s; l  r * s - Suc 0 |] ==> rtos r s l  s * r - Suc 0

lemma rtos_hom2:

  [| 0 < r; 0 < s; l  r * s - Suc 0 |] ==> rtos r s l  r * s - Suc 0

lemma rtos_hom3:

  [| 0 < r; 0 < s; i  r * s - Suc 0 |] ==> rtos r s i div r = i mod s

lemma rtos_hom3_1:

  [| 0 < r; 0 < s; i  r * s - Suc 0 |] ==> rtos r s i mod r = i div s

lemma rtos_hom5:

  [| 0 < r; 0 < s; i  r * s - Suc 0; i div s = r - Suc 0 |]
  ==> Suc (rtos r s i) div r = Suc (i mod s)

lemma rtos_hom7:

  [| 0 < r; 0 < s; i  r * s - Suc 0; i div s = r - Suc 0 |]
  ==> Suc (rtos r s i) mod r = 0

lemma rtos_inj:

  [| 0 < r; 0 < s |] ==> inj_on (rtos r s) {i. i  r * s - Suc 0}

lemma rtos_rs_Tr1:

  [| 0 < r; 0 < s |] ==> rtos r s (r * s) = r * s

lemma rtos_rs_Tr2:

  [| 0 < r; 0 < s |] ==> ∀lr * s. rtos r s l  r * s

lemma rtos_rs_Tr3:

  [| 0 < r; 0 < s |] ==> inj_on (rtos r s) {i. i  r * s}

lemma Qw_cmpser:

  [| Group G; w_cmpser G (Suc n) f |] ==> Gchain n (Qw_cmpser G f)

lemma rfn_tool8:

  [| compseries G r f; 0 < r |] ==> d_gchain G r f

lemma rfn_tool16:

  [| 0 < r; 0 < s; i  s * r - Suc 0; G » f (i div s) ;
     \<natural>f (i div s) \<triangleright> f (Suc (i div s));
     \<natural>f (i div s) » f (i div s) ∩ g (s - Suc 0)  |]
  ==> \<natural>f (Suc (i div
                        s)) \<struct>G f (i div s) ∩
                                       g (s -
  Suc 0) \<triangleright> f (Suc (i div s))

lemma rfn_tool30:

  [| 0 < r; 0 < s; l div s * s + s < s * r |] ==> Suc (l div s) < r

lemma simple_grouptr0:

  [| G » H ; G \<triangleright> K; K  H; simple_Group (G / K) |]
  ==> H = carrier GH = K

lemma compser_nsg:

  [| 0 < n; compseries G n f; i  n - 1 |]
  ==> \<natural>f i \<triangleright> f (Suc i)

lemma compseriesTr5:

  [| 0 < n; compseries G n f; i  n - Suc 0 |] ==> f (Suc i)  f i

lemma refine_cmpserTr0:

  [| 0 < n; compseries G n f; i  n - 1; G » H ; f (Suc i)  HH  f i |]
  ==> H = f (Suc i) ∨ H = f i

lemma div_Tr4:

  [| 0 < r; 0 < s; j < s * r |] ==> j div s * s + s  r * s

lemma compseries_is_tW_cmpser:

  [| 0 < r; compseries G r f |] ==> tW_cmpser G r f

lemma compseries_is_td_gchain:

  [| 0 < r; compseries G r f |] ==> td_gchain G r f

lemma compseries_is_D_gchain:

  [| 0 < r; compseries G r f |] ==> D_gchain G r f

lemma divTr5:

  [| 0 < r; 0 < s; l < r * s |] ==> l div s * s  ll  Suc (l div s) * s

lemma rfn_compseries_iMTr1:

  [| 0 < r; 0 < s; compseries G r f; h ∈ wcsr_rfns G r f s |]
  ==> f ` {i. i  r}  h ` {i. i  s * r}

lemma rfn_compseries_iMTr2:

  [| 0 < r; 0 < s; xa < s * r |]
  ==> xa div s * s  r * s ∧ Suc (xa div s) * s  r * s

lemma rfn_compseries_iMTr3:

  [| 0 < r; 0 < s; compseries G r f; j  r; ∀ir. h (i * s) = f i |]
  ==> h (j * s) = f j

lemma rfn_compseries_iM:

  [| 0 < r; 0 < s; compseries G r f; h ∈ wcsr_rfns G r f s |]
  ==> card (h ` {i. i  s * r}) = r + 1

lemma cmp_rfn0:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g; i  r - 1; j  s - 1 |]
  ==> G » f (Suc i) \<struct>G f ig j 

lemma cmp_rfn1:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g |]
  ==> f (Suc 0) \<struct>G f 0g 0 = carrier G

lemma cmp_rfn2:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g; l  s * r |]
  ==> G » cmp_rfn G r f s g l 

lemma cmp_rfn3:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g |]
  ==> cmp_rfn G r f s g 0 = carrier G ∧ cmp_rfn G r f s g (s * r) = {\<one>}

lemma rfn_tool20:

  [| 0 < m; a = b * m + c; c < m |] ==> a mod m = c

lemma Suci_mod_s_2:

  [| 0 < r; 0 < s; i  r * s - Suc 0; i mod s < s - Suc 0 |]
  ==> Suc i mod s = Suc (i mod s)

lemma inter_sgsTr1:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g; i < r * s |]
  ==> G » f (i div s) ∩ g (s - Suc 0) 

lemma JHS_Tr0_2:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g |]
  ==> ∀is * r - Suc 0.
         \<natural>cmp_rfn G r f s g i \<triangleright> cmp_rfn G r f s g (Suc i)

lemma cmp_rfn4:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g; l  s * r - Suc 0 |]
  ==> cmp_rfn G r f s g (Suc l)  cmp_rfn G r f s g l

lemma cmp_rfn5:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g |]
  ==> ∀ir. cmp_rfn G r f s g (i * s) = f i

lemma JHS_Tr0:

  [| 0 < r; 0 < s; compseries G r f; compseries G s g |]
  ==> cmp_rfn G r f s g ∈ wcsr_rfns G r f s

lemma rfn_tool17:

  a = b ==> a - c = b - c

lemma isom4b:

  [| Group G; G \<triangleright> N; G » H  |]
  ==> (\<natural>GN \<struct>G H / N) ≅ (\<natural>GH / (HN))

lemma Suc_rtos_div_r_1:

  [| 0 < r; 0 < s; i  r * s - Suc 0; Suc (rtos r s i) < r * s;
     i mod s = s - Suc 0; i div s < r - Suc 0 |]
  ==> Suc (rtos r s i) div r = i mod s

lemma Suc_rtos_mod_r_1:

  [| 0 < r; 0 < s; i  r * s - Suc 0; Suc (rtos r s i) < r * s;
     i mod s = s - Suc 0; i div s < r - Suc 0 |]
  ==> Suc (rtos r s i) mod r = Suc (i div s)

lemma i_div_s_less:

  [| 0 < r; 0 < s; i  r * s - Suc 0; Suc (rtos r s i) < r * s;
     i mod s = s - Suc 0; Suc i < s * r |]
  ==> i div s < r - Suc 0

lemma rtos_mod_r_1:

  [| 0 < r; 0 < s; i  r * s - Suc 0; rtos r s i < r * s; i mod s = s - Suc 0 |]
  ==> rtos r s i mod r = i div s

lemma Suc_i_mod_s_0_1:

  [| 0 < r; 0 < s; i  r * s - Suc 0; i mod s = s - Suc 0 |] ==> Suc i mod s = 0

lemma Suci_div_s_2:

  [| 0 < r; 0 < s; i  r * s - Suc 0; i mod s < s - Suc 0 |]
  ==> Suc i div s = i div s

lemma rtos_i_mod_r_2:

  [| 0 < r; 0 < s; i  r * s - Suc 0 |] ==> rtos r s i mod r = i div s

lemma Suc_rtos_i_mod_r_2:

  [| 0 < r; 0 < s; i  r * s - Suc 0; i div s = r - Suc 0 |]
  ==> Suc (rtos r s i) mod r = 0

lemma Suc_rtos_i_mod_r_3:

  [| 0 < r; 0 < s; i  r * s - Suc 0; i div s < r - Suc 0 |]
  ==> Suc (rtos r s i) mod r = Suc (i div s)

lemma Suc_rtos_div_r_3:

  [| 0 < r; 0 < s; i mod s < s - Suc 0; i  r * s - Suc 0;
     Suc (rtos r s i) < r * s; i div s < r - Suc 0 |]
  ==> Suc (rtos r s i) div r = i mod s

lemma r_s_div_s:

  [| 0 < r; 0 < s |] ==> (r * s - Suc 0) div s = r - Suc 0

lemma r_s_mod_s:

  [| 0 < r; 0 < s |] ==> (r * s - Suc 0) mod s = s - Suc 0

lemma rtos_r_s:

  [| 0 < r; 0 < s |] ==> rtos r s (r * s - Suc 0) = r * s - Suc 0

lemma rtos_rs_1:

  [| 0 < r; 0 < s; rtos r s i < r * s; ¬ Suc (rtos r s i) < r * s |]
  ==> rtos r s i = r * s - Suc 0

lemma rtos_rs_i_rs:

  [| 0 < r; 0 < s; i  r * s - Suc 0; rtos r s i = r * s - Suc 0 |]
  ==> i = r * s - Suc 0

lemma JHS_Tr1_1:

  [| Group G; 0 < r; 0 < s; compseries G r f; compseries G s g |]
  ==> f (Suc ((r * s - Suc 0) div
              s)) \<struct>G f ((r * s - Suc 0) div s) ∩
                             g ((r * s - Suc 0) mod s) =
      f (r - Suc 0) ∩ g (s - Suc 0)

lemma JHS_Tr1_2:

  [| Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; k < r - Suc 0 |]
  ==> (\<natural>Gf (Suc k) \<struct>G f kg (s - Suc 0) /
       (f (Suc (Suc k)) \<struct>G f (Suc k) ∩ g 0)) ≅
      (\<natural>Gg s \<struct>G g (s - Suc 0) ∩ f k /
       (g s \<struct>G g (s - Suc 0) ∩ f (Suc k)))

lemma JHS_Tr1_3:

  [| Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; i  s * r - Suc 0;
     Suc (rtos r s i) < s * r; Suc i < s * r; i mod s < s - Suc 0;
     Suc i div s  r - Suc 0; i div s = r - Suc 0 |]
  ==> Group
       (\<natural>Gf r \<struct>G f (r - Suc 0) ∩ g (i mod s) /
        (f r \<struct>G f (r - Suc 0) ∩ g (Suc (i mod s))))

lemma JHS_Tr1_4:

  [| Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; i  s * r - Suc 0;
     Suc (rtos r s i) < s * r; Suc i < s * r; i mod s < s - Suc 0;
     Suc i div s  r - Suc 0; i div s = r - Suc 0 |]
  ==> Group
       (\<natural>Gg (Suc (i mod s)) \<struct>G g (i mod s) ∩ f (r - Suc 0) /
        (g (Suc (Suc (i mod s))) \<struct>G g (Suc (i mod s)) ∩ f 0))

lemma JHS_Tr1_5:

  [| Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; i  s * r - Suc 0;
     Suc (rtos r s i) < s * r; Suc i < s * r; i mod s < s - Suc 0;
     i div s < r - Suc 0 |]
  ==> (\<natural>Gf (Suc (i div s)) \<struct>G f (i div s) ∩ g (i mod s) /
       (f (Suc (i div s)) \<struct>G f (i div s) ∩ g (Suc (i mod s)))) ≅
      (\<natural>Gg (Suc (i mod s)) \<struct>G g (i mod s) ∩ f (i div s) /
       (g (Suc (Suc (rtos r s i) div
                r)) \<struct>G g (Suc (rtos r s i) div r) ∩
                               f (Suc (rtos r s i) mod r)))

lemma JHS_Tr1_6:

  [| Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; i  r * s - Suc 0;
     Suc (rtos r s i) < r * s |]
  ==> (\<natural>Gcmp_rfn G r f s g i / cmp_rfn G r f s g (Suc i)) ≅
      (\<natural>Gg (Suc (rtos r s i div
                          r)) \<struct>G g (rtos r s i div r) ∩
                                         f (rtos r s i mod r) /
       (g (Suc (Suc (rtos r s i) div
                r)) \<struct>G g (Suc (rtos r s i) div r) ∩
                               f (Suc (rtos r s i) mod r)))

lemma JHS_Tr1:

  [| Group G; 0 < r; 0 < s; compseries G r f; compseries G s g |]
  ==> isom_Gchains (r * s - 1) (rtos r s) (Qw_cmpser G (cmp_rfn G r f s g))
       (Qw_cmpser G (cmp_rfn G s g r f))

lemma abc_SucTr0:

  [| 0 < a; c  b; a - Suc 0 = b - c |] ==> a = Suc b - c

lemma length_wcmpser0_0:

  [| Group G; Ugp E; w_cmpser G (Suc 0) f |]
  ==> f ` {i. i  Suc 0} = {f 0, f (Suc 0)}

lemma length_wcmpser0_1:

  [| Group G; Ugp E; w_cmpser G (Suc n) f; i ∈ {i. i  n}; Qw_cmpser G f iE |]
  ==> f i = f (Suc i)

lemma length_wcmpser0_2:

  [| Group G; Ugp E; w_cmpser G (Suc n) f; i  n; ¬ Qw_cmpser G f iE |]
  ==> f i  f (Suc i)

lemma length_wcmpser0_3:

  [| Group G; Ugp E; w_cmpser G (Suc (Suc n)) f; f (Suc n)  f (Suc (Suc n)) |]
  ==> f (Suc (Suc n))  f ` {i. i  Suc n}

lemma length_wcmpser0_4:

  [| Group G; Ugp E; w_cmpser G (Suc 0) f |]
  ==> card (f ` {i. i  Suc 0}) - 1 =
      Suc 0 - card {i. i = 0 ∧ Qw_cmpser G f iE}

lemma length_wcmpser0_5:

  [| Group G; Ugp E; w_cmpser G (Suc (Suc n)) f; w_cmpser G (Suc n) f;
     card (f ` {i. i  Suc n}) - 1 =
     Suc n - card {i. i  n ∧ Qw_cmpser G f iE};
     Qw_cmpser G f (Suc n) ≅ E |]
  ==> card (f ` {i. i  Suc (Suc n)}) - 1 =
      Suc (Suc n) - card {i. i  Suc n ∧ Qw_cmpser G f iE}

lemma length_wcmpser0_6:

  [| Group G; w_cmpser G (Suc (Suc n)) f |] ==> 0 < card (f ` {i. i  Suc n})

lemma length_wcmpser0_7:

  [| Group G; w_cmpser G (Suc (Suc n)) f |]
  ==> card {i. i  n ∧ Qw_cmpser G f iE}  Suc n

lemma length_wcmpser0:

  [| Group G; Ugp E |]
  ==> ∀f. w_cmpser G (Suc n) f -->
          card (f ` {i. i  Suc n}) - 1 =
          Suc n - card {i. i  n ∧ Qw_cmpser G f iE}

lemma length_of_twcmpser:

  [| Group G; Ugp E; tw_cmpser G (Suc n) f |]
  ==> length_twcmpser G (Suc n) f = Suc n - card {i. i  n ∧ Qw_cmpser G f iE}

lemma JHS_1:

  [| Group G; Ugp E; compseries G r f; compseries G s g; 0 < r; 0 < s |]
  ==> r = r * s -
          card {i. i  r * s - Suc 0 ∧ Qw_cmpser G (cmp_rfn G r f s g) iE}

lemma J_H_S:

  [| Group G; Ugp E; compseries G r f; compseries G s g; 0 < r; 0 < s |] ==> r = s