src/HOL/Algebra/Generated_Groups.thy
author desharna
Thu, 08 Jul 2021 08:42:36 +0200
changeset 73932 fd21b4a93043
parent 70039 733e256ecdf3
permissions -rw-r--r--
added opaque_combs and renamed hide_lams to opaque_lifting
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68582
b9b9e2985878 more standard headers;
wenzelm
parents: 68576
diff changeset
     1
(*  Title:      HOL/Algebra/Generated_Groups.thy
b9b9e2985878 more standard headers;
wenzelm
parents: 68576
diff changeset
     2
    Author:     Paulo Emílio de Vilhena
b9b9e2985878 more standard headers;
wenzelm
parents: 68576
diff changeset
     3
*)
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
     5
section \<open>Generated Groups\<close>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
     6
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory Generated_Groups
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
  imports Group Coset
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
     9
  
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
begin
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
    12
subsection \<open>Generated Groups\<close>
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    14
inductive_set generate :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set"
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
  for G and H where
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
    one:  "\<one>\<^bsub>G\<^esub> \<in> generate G H"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  | incl: "h \<in> H \<Longrightarrow> h \<in> generate G H"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
  | inv:  "h \<in> H \<Longrightarrow> inv\<^bsub>G\<^esub> h \<in> generate G H"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
  | eng:  "h1 \<in> generate G H \<Longrightarrow> h2 \<in> generate G H \<Longrightarrow> h1 \<otimes>\<^bsub>G\<^esub> h2 \<in> generate G H"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
    22
subsubsection \<open>Basic Properties\<close>
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    24
lemma (in group) generate_consistent:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    25
  assumes "K \<subseteq> H" "subgroup H G" shows "generate (G \<lparr> carrier := H \<rparr>) K = generate G K"
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
proof
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    27
  show "generate (G \<lparr> carrier := H \<rparr>) K \<subseteq> generate G K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    28
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    29
    fix h assume "h \<in> generate (G \<lparr> carrier := H \<rparr>) K" thus "h \<in> generate G K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    30
    proof (induction, simp add: one, simp_all add: incl[of _ K G] eng)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    31
      case inv thus ?case
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    32
        using m_inv_consistent assms generate.inv[of _ K G] by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    33
    qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    34
  qed
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
next
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    36
  show "generate G K \<subseteq> generate (G \<lparr> carrier := H \<rparr>) K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    37
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    38
    note gen_simps = one incl eng
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    39
    fix h assume "h \<in> generate G K" thus "h \<in> generate (G \<lparr> carrier := H \<rparr>) K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    40
      using gen_simps[where ?G = "G \<lparr> carrier := H \<rparr>"]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    41
    proof (induction, auto)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    42
      fix h assume "h \<in> K" thus "inv h \<in> generate (G \<lparr> carrier := H \<rparr>) K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    43
        using m_inv_consistent assms generate.inv[of h K "G \<lparr> carrier := H \<rparr>"] by auto
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    48
lemma (in group) generate_in_carrier:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    49
  assumes "H \<subseteq> carrier G" and "h \<in> generate G H" shows "h \<in> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    50
  using assms(2,1) by (induct h rule: generate.induct) (auto)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    51
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    52
lemma (in group) generate_incl:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    53
  assumes "H \<subseteq> carrier G" shows "generate G H \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    54
  using generate_in_carrier[OF assms(1)] by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    55
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    56
lemma (in group) generate_m_inv_closed:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    57
  assumes "H \<subseteq> carrier G" and "h \<in> generate G H" shows "(inv h) \<in> generate G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    58
  using assms(2,1)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    59
proof (induction rule: generate.induct, auto simp add: one inv incl)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    60
  fix h1 h2
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    61
  assume h1: "h1 \<in> generate G H" "inv h1 \<in> generate G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    62
     and h2: "h2 \<in> generate G H" "inv h2 \<in> generate G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    63
  hence "inv (h1 \<otimes> h2) = (inv h2) \<otimes> (inv h1)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    64
    by (meson assms generate_in_carrier group.inv_mult_group is_group)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    65
  thus "inv (h1 \<otimes> h2) \<in> generate G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    66
    using generate.eng[OF h2(2) h1(2)] by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    67
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    68
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    69
lemma (in group) generate_is_subgroup:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    70
  assumes "H \<subseteq> carrier G" shows "subgroup (generate G H) G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    71
  using subgroup.intro[OF generate_incl eng one generate_m_inv_closed] assms by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    72
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    73
lemma (in group) mono_generate:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    74
  assumes "K \<subseteq> H" shows "generate G K \<subseteq> generate G H"
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
proof
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    76
  fix h assume "h \<in> generate G K" thus "h \<in> generate G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    77
    using assms by (induction) (auto simp add: one incl inv eng)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    78
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    79
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    80
lemma (in group) generate_subgroup_incl:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    81
  assumes "K \<subseteq> H" "subgroup H G" shows "generate G K \<subseteq> H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    82
  using group.generate_incl[OF subgroup_imp_group[OF assms(2)], of K] assms(1)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    83
  by (simp add: generate_consistent[OF assms])
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    84
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    85
lemma (in group) generate_minimal:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    86
  assumes "H \<subseteq> carrier G" shows "generate G H = \<Inter> { H'. subgroup H' G \<and> H \<subseteq> H' }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    87
  using generate_subgroup_incl generate_is_subgroup[OF assms] incl[of _ H] by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    88
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    89
lemma (in group) generateI:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    90
  assumes "subgroup E G" "H \<subseteq> E" and "\<And>K. \<lbrakk> subgroup K G; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    91
  shows "E = generate G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    92
proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    93
  have subset: "H \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    94
    using subgroup.subset assms by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    95
  show ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    96
    using assms unfolding generate_minimal[OF subset] by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    97
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    98
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    99
lemma (in group) normal_generateI:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   100
  assumes "H \<subseteq> carrier G" and "\<And>h g. \<lbrakk> h \<in> H; g \<in> carrier G \<rbrakk> \<Longrightarrow> g \<otimes> h \<otimes> (inv g) \<in> H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   101
  shows "generate G H \<lhd> G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   102
proof (rule normal_invI[OF generate_is_subgroup[OF assms(1)]])
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   103
  fix g h assume g: "g \<in> carrier G" show "h \<in> generate G H \<Longrightarrow> g \<otimes> h \<otimes> (inv g) \<in> generate G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   104
  proof (induct h rule: generate.induct)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   105
    case one thus ?case
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   106
      using g generate.one by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   107
  next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   108
    case incl show ?case
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   109
      using generate.incl[OF assms(2)[OF incl g]] .
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   110
  next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   111
    case (inv h)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   112
    hence h: "h \<in> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   113
      using assms(1) by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   114
    hence "inv (g \<otimes> h \<otimes> (inv g)) = g \<otimes> (inv h) \<otimes> (inv g)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   115
      using g by (simp add: inv_mult_group m_assoc)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   116
    thus ?case
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   117
      using generate_m_inv_closed[OF assms(1) generate.incl[OF assms(2)[OF inv g]]] by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   118
  next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   119
    case (eng h1 h2)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   120
    note in_carrier = eng(1,3)[THEN generate_in_carrier[OF assms(1)]]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   121
    have "g \<otimes> (h1 \<otimes> h2) \<otimes> inv g = (g \<otimes> h1 \<otimes> inv g) \<otimes> (g \<otimes> h2 \<otimes> inv g)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   122
      using in_carrier g by (simp add: inv_solve_left m_assoc)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   123
    thus ?case
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   124
      using generate.eng[OF eng(2,4)] by simp
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   128
lemma (in group) subgroup_int_pow_closed:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   129
  assumes "subgroup H G" "h \<in> H" shows "h [^] (k :: int) \<in> H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   130
  using group.int_pow_closed[OF subgroup_imp_group[OF assms(1)]] assms(2)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   131
  unfolding int_pow_consistent[OF assms] by simp
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   133
lemma (in group) generate_pow:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   134
  assumes "a \<in> carrier G" shows "generate G { a } = { a [^] (k :: int) | k. k \<in> UNIV }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   135
proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   136
  show "{ a [^] (k :: int) | k. k \<in> UNIV } \<subseteq> generate G { a }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   137
    using subgroup_int_pow_closed[OF generate_is_subgroup[of "{ a }"] incl[of a]] assms by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   138
next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   139
  show "generate G { a } \<subseteq> { a [^] (k :: int) | k. k \<in> UNIV }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   140
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   141
    fix h assume "h \<in> generate G { a }" hence "\<exists>k :: int. h = a [^] k"
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   142
    proof (induction)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   143
      case one
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   144
      then show ?case
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   145
        using int_pow_0 [of G] by metis
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   146
    next
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   147
      case (incl h)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   148
      then show ?case
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   149
        by (metis assms int_pow_1 singletonD)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   150
    next
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   151
      case (inv h)
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   152
      then show ?case
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   153
        by (metis assms int_pow_1 int_pow_neg singletonD)
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
    next
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   155
      case (eng h1 h2)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   156
      then show ?case
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   157
        using assms by (metis int_pow_mult)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   158
    qed
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   159
    then show "h \<in> { a [^] (k :: int) | k. k \<in> UNIV }"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   160
      by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   161
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   162
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   163
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   164
corollary (in group) generate_one: "generate G { \<one> } = { \<one> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   165
  using generate_pow[of "\<one>", OF one_closed] by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   166
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   167
corollary (in group) generate_empty: "generate G {} = { \<one> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   168
  using mono_generate[of "{}" "{ \<one> }"] generate.one unfolding generate_one by auto
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   170
lemma (in group_hom)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   171
  "subgroup K G \<Longrightarrow> subgroup (h ` K) H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   172
  using subgroup_img_is_subgroup by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   173
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   174
lemma (in group_hom) generate_img:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   175
  assumes "K \<subseteq> carrier G" shows "generate H (h ` K) = h ` (generate G K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   176
proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   177
  have "h ` K \<subseteq> h ` (generate G K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   178
    using incl[of _ K G] by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   179
  thus "generate H (h ` K) \<subseteq> h ` (generate G K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   180
    using generate_subgroup_incl subgroup_img_is_subgroup[OF G.generate_is_subgroup[OF assms]] by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   181
next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   182
  show "h ` (generate G K) \<subseteq> generate H (h ` K)"
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
  proof
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   184
    fix a assume "a \<in> h ` (generate G K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   185
    then obtain k where "k \<in> generate G K" "a = h k"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   186
      by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   187
    show "a \<in> generate H (h ` K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   188
      using \<open>k \<in> generate G K\<close> unfolding \<open>a = h k\<close>
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   189
    proof (induct k, auto simp add: generate.one[of H] generate.incl[of _ "h ` K" H])
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   190
      case (inv k) show ?case
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   191
        using assms generate.inv[of "h k" "h ` K" H] inv by auto  
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   192
    next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   193
      case eng show ?case
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   194
        using generate.eng[OF eng(2,4)] eng(1,3)[THEN G.generate_in_carrier[OF assms]] by auto
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   200
subsection \<open>Derived Subgroup\<close>
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   202
subsubsection \<open>Definitions\<close>
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   204
abbreviation derived_set :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   205
  where "derived_set G H \<equiv>
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   206
           \<Union>h1 \<in> H. (\<Union>h2 \<in> H. { h1 \<otimes>\<^bsub>G\<^esub> h2 \<otimes>\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> h1) \<otimes>\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> h2) })"
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
definition derived :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set" where
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
  "derived G H = generate G (derived_set G H)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   211
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   212
subsubsection \<open>Basic Properties\<close>
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   213
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
lemma (in group) derived_set_incl:
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   215
  assumes "K \<subseteq> H" "subgroup H G" shows "derived_set G K \<subseteq> H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   216
  using assms(1) subgroupE(3-4)[OF assms(2)] by (auto simp add: subset_iff)
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
lemma (in group) derived_incl:
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   219
  assumes "K \<subseteq> H" "subgroup H G" shows "derived G K \<subseteq> H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   220
  using generate_subgroup_incl[OF derived_set_incl] assms unfolding derived_def by auto
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
lemma (in group) derived_set_in_carrier:
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   223
  assumes "H \<subseteq> carrier G" shows "derived_set G H \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   224
  using derived_set_incl[OF assms subgroup_self] .
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   225
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   226
lemma (in group) derived_in_carrier:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   227
  assumes "H \<subseteq> carrier G" shows "derived G H \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   228
  using derived_incl[OF assms subgroup_self] .
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   229
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   230
lemma (in group) exp_of_derived_in_carrier:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   231
  assumes "H \<subseteq> carrier G" shows "(derived G ^^ n) H \<subseteq> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   232
  using assms derived_in_carrier by (induct n) (auto)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   233
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   234
lemma (in group) derived_is_subgroup:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   235
  assumes "H \<subseteq> carrier G" shows "subgroup (derived G H) G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   236
  unfolding derived_def using generate_is_subgroup[OF derived_set_in_carrier[OF assms]] .
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   237
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   238
lemma (in group) exp_of_derived_is_subgroup:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   239
  assumes "subgroup H G" shows "subgroup ((derived G ^^ n) H) G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   240
  using assms derived_is_subgroup subgroup.subset by (induct n) (auto)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   241
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   242
lemma (in group) exp_of_derived_is_subgroup':
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   243
  assumes "H \<subseteq> carrier G" shows "subgroup ((derived G ^^ (Suc n)) H) G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   244
  using assms derived_is_subgroup[OF subgroup.subset] derived_is_subgroup by (induct n) (auto)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   245
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   246
lemma (in group) mono_derived_set:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   247
  assumes "K \<subseteq> H" shows "derived_set G K \<subseteq> derived_set G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   248
  using assms by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   249
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   250
lemma (in group) mono_derived:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   251
  assumes "K \<subseteq> H" shows "derived G K \<subseteq> derived G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   252
  unfolding derived_def using mono_generate[OF mono_derived_set[OF assms]] .
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   253
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   254
lemma (in group) mono_exp_of_derived:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   255
  assumes "K \<subseteq> H" shows "(derived G ^^ n) K \<subseteq> (derived G ^^ n) H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   256
  using assms mono_derived by (induct n) (auto)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   257
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   258
lemma (in group) derived_set_consistent:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   259
  assumes "K \<subseteq> H" "subgroup H G" shows "derived_set (G \<lparr> carrier := H \<rparr>) K = derived_set G K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   260
  using m_inv_consistent[OF assms(2)] assms(1) by (auto simp add: subset_iff)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   261
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   262
lemma (in group) derived_consistent:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   263
  assumes "K \<subseteq> H" "subgroup H G" shows "derived (G \<lparr> carrier := H \<rparr>) K = derived G K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   264
  using generate_consistent[OF derived_set_incl] derived_set_consistent assms by (simp add: derived_def)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   265
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   266
lemma (in comm_group) derived_eq_singleton:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   267
  assumes "H \<subseteq> carrier G" shows "derived G H = { \<one> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   268
proof (cases "derived_set G H = {}")
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   269
  case True show ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   270
    using generate_empty unfolding derived_def True by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   271
next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   272
  case False
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   273
  have aux_lemma: "h \<in> derived_set G H \<Longrightarrow> h = \<one>" for h
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   274
    using assms by (auto simp add: subset_iff)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   275
       (metis (no_types, lifting) m_comm m_closed inv_closed inv_solve_right l_inv l_inv_ex)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   276
  have "derived_set G H = { \<one> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   277
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   278
    show "derived_set G H \<subseteq> { \<one> }"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   279
      using aux_lemma by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   280
  next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   281
    obtain h where h: "h \<in> derived_set G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   282
      using False by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   283
    thus "{ \<one> } \<subseteq> derived_set G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   284
      using aux_lemma[OF h] by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   285
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   286
  thus ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   287
    using generate_one unfolding derived_def by auto
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
lemma (in group) derived_is_normal:
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   291
  assumes "H \<lhd> G" shows "derived G H \<lhd> G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   292
proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   293
  interpret H: normal H G
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   294
    using assms .
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   295
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   296
  show ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   297
    unfolding derived_def
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   298
  proof (rule normal_generateI[OF derived_set_in_carrier[OF H.subset]])
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
    fix h g assume "h \<in> derived_set G H" and g: "g \<in> carrier G"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   300
    then obtain h1 h2 where h: "h1 \<in> H" "h2 \<in> H" "h = h1 \<otimes> h2 \<otimes> inv h1 \<otimes> inv h2"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   301
      by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   302
    hence in_carrier: "h1 \<in> carrier G" "h2 \<in> carrier G" "g \<in> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   303
      using H.subset g by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   304
    have "g \<otimes> h \<otimes> inv g =
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
           g \<otimes> h1 \<otimes> (inv g \<otimes> g) \<otimes> h2 \<otimes> (inv g \<otimes> g) \<otimes> inv h1 \<otimes> (inv g \<otimes> g) \<otimes> inv h2 \<otimes> inv g"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   306
      unfolding h(3) by (simp add: in_carrier m_assoc)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   307
    also have " ... =
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
          (g \<otimes> h1 \<otimes> inv g) \<otimes> (g \<otimes> h2 \<otimes> inv g) \<otimes> (g \<otimes> inv h1 \<otimes> inv g) \<otimes> (g \<otimes> inv h2 \<otimes> inv g)"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   309
      using in_carrier m_assoc inv_closed m_closed by presburger
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   310
    finally have "g \<otimes> h \<otimes> inv g =
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   311
          (g \<otimes> h1 \<otimes> inv g) \<otimes> (g \<otimes> h2 \<otimes> inv g) \<otimes> inv (g \<otimes> h1 \<otimes> inv g) \<otimes> inv (g \<otimes> h2 \<otimes> inv g)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   312
      by (simp add: in_carrier inv_mult_group m_assoc)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   313
    thus "g \<otimes> h \<otimes> inv g \<in> derived_set G H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   314
      using h(1-2)[THEN H.inv_op_closed2[OF g]] by auto
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   318
lemma (in group) normal_self: "carrier G \<lhd> G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   319
  by (rule normal_invI[OF subgroup_self], simp)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   320
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
corollary (in group) derived_self_is_normal: "derived G (carrier G) \<lhd> G"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   322
  using derived_is_normal[OF normal_self] .
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
corollary (in group) derived_subgroup_is_normal:
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   325
  assumes "subgroup H G" shows "derived G H \<lhd> G \<lparr> carrier := H \<rparr>"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   326
  using group.derived_self_is_normal[OF subgroup_imp_group[OF assms]]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   327
        derived_consistent[OF _ assms]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   328
  by simp
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
corollary (in group) derived_quot_is_group: "group (G Mod (derived G (carrier G)))"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   331
  using normal.factorgroup_is_group[OF derived_self_is_normal] by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   332
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   333
lemma (in group) derived_quot_is_comm_group: "comm_group (G Mod (derived G (carrier G)))"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   334
proof (rule group.group_comm_groupI[OF derived_quot_is_group], simp add: FactGroup_def)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   335
  interpret DG: normal "derived G (carrier G)" G
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   336
    using derived_self_is_normal .
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   338
  fix H K assume "H \<in> rcosets derived G (carrier G)" and "K \<in> rcosets derived G (carrier G)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   339
  then obtain g1 g2
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   340
    where g1: "g1 \<in> carrier G" "H = derived G (carrier G) #> g1"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   341
      and g2: "g2 \<in> carrier G" "K = derived G (carrier G) #> g2"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   342
    unfolding RCOSETS_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   343
  hence "H <#> K = derived G (carrier G) #> (g1 \<otimes> g2)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   344
    by (simp add: DG.rcos_sum)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   345
  also have " ... = derived G (carrier G) #> (g2 \<otimes> g1)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   346
  proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   347
    { fix g1 g2 assume g1: "g1 \<in> carrier G" and g2: "g2 \<in> carrier G"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   348
      have "derived G (carrier G) #> (g1 \<otimes> g2) \<subseteq> derived G (carrier G) #> (g2 \<otimes> g1)"
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
      proof
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   350
        fix h assume "h \<in> derived G (carrier G) #> (g1 \<otimes> g2)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   351
        then obtain g' where h: "g' \<in> carrier G" "g' \<in> derived G (carrier G)" "h = g' \<otimes> (g1 \<otimes> g2)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   352
          using DG.subset unfolding r_coset_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   353
        hence "h = g' \<otimes> (g1 \<otimes> g2) \<otimes> (inv g1 \<otimes> inv g2 \<otimes> g2 \<otimes> g1)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   354
          using g1 g2 by (simp add: m_assoc)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   355
        hence "h = (g' \<otimes> (g1 \<otimes> g2 \<otimes> inv g1 \<otimes> inv g2)) \<otimes> (g2 \<otimes> g1)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   356
          using h(1) g1 g2 inv_closed m_assoc m_closed by presburger
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   357
        moreover have "g1 \<otimes> g2 \<otimes> inv g1 \<otimes> inv g2 \<in> derived G (carrier G)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   358
          using incl[of _ "derived_set G (carrier G)"] g1 g2 unfolding derived_def by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   359
        hence "g' \<otimes> (g1 \<otimes> g2 \<otimes> inv g1 \<otimes> inv g2) \<in> derived G (carrier G)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   360
          using DG.m_closed[OF h(2)] by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   361
        ultimately show "h \<in> derived G (carrier G) #> (g2 \<otimes> g1)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   362
          unfolding r_coset_def by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   363
      qed }
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   364
    thus ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   365
      using g1(1) g2(1) by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   366
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   367
  also have " ... = K <#> H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   368
    by (simp add: g1 g2 DG.rcos_sum)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   369
  finally show "H <#> K = K <#> H" .
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
corollary (in group) derived_quot_of_subgroup_is_comm_group:
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   373
  assumes "subgroup H G" shows "comm_group ((G \<lparr> carrier := H \<rparr>) Mod (derived G H))"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   374
  using group.derived_quot_is_comm_group[OF subgroup_imp_group[OF assms]]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   375
        derived_consistent[OF _ assms]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   376
  by simp
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   378
proposition (in group) derived_minimal:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   379
  assumes "H \<lhd> G" and "comm_group (G Mod H)" shows "derived G (carrier G) \<subseteq> H"
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
proof -
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   381
  interpret H: normal H G
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   382
    using assms(1) .
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
  show ?thesis
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   385
    unfolding derived_def
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   386
  proof (rule generate_subgroup_incl[OF _ H.subgroup_axioms])
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   387
    show "derived_set G (carrier G) \<subseteq> H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   388
    proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   389
      fix h assume "h \<in> derived_set G (carrier G)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   390
      then obtain g1 g2 where h: "g1 \<in> carrier G" "g2 \<in> carrier G" "h = g1 \<otimes> g2 \<otimes> inv g1 \<otimes> inv g2"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   391
        by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   392
      have "H #> (g1 \<otimes> g2) = (H #> g1) <#> (H #> g2)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   393
        by (simp add: h(1-2) H.rcos_sum)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   394
      also have " ... = (H #> g2) <#> (H #> g1)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   395
        using comm_groupE(4)[OF assms(2)] h(1-2) unfolding FactGroup_def RCOSETS_def by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   396
      also have " ... = H #> (g2 \<otimes> g1)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   397
        by (simp add: h(1-2) H.rcos_sum)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   398
      finally have "H #> (g1 \<otimes> g2) = H #> (g2 \<otimes> g1)" .
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   399
      then obtain h' where "h' \<in> H" "\<one> \<otimes> (g1 \<otimes> g2) = h' \<otimes> (g2 \<otimes> g1)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   400
        using H.one_closed unfolding r_coset_def by blast
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   401
      thus "h \<in> H"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   402
        using h m_assoc by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   403
    qed
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   407
proposition (in group) derived_of_subgroup_minimal:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   408
  assumes "K \<lhd> G \<lparr> carrier := H \<rparr>" "subgroup H G" and "comm_group ((G \<lparr> carrier := H \<rparr>) Mod K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   409
  shows "derived G H \<subseteq> K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   410
  using group.derived_minimal[OF subgroup_imp_group[OF assms(2)] assms(1,3)]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   411
        derived_consistent[OF _ assms(2)]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   412
  by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   413
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   414
lemma (in group_hom) derived_img:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   415
  assumes "K \<subseteq> carrier G" shows "derived H (h ` K) = h ` (derived G K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   416
proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   417
  have "derived_set H (h ` K) = h ` (derived_set G K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   418
  proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   419
    show "derived_set H (h ` K) \<subseteq> h ` derived_set G K"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   420
    proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   421
      fix a assume "a \<in> derived_set H (h ` K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   422
      then obtain k1 k2
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   423
        where "k1 \<in> K" "k2 \<in> K" "a = (h k1) \<otimes>\<^bsub>H\<^esub> (h k2) \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k1) \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k2)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   424
        by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   425
      hence "a = h (k1 \<otimes> k2 \<otimes> inv k1 \<otimes> inv k2)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   426
        using assms by (simp add: subset_iff)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   427
      from this \<open>k1 \<in> K\<close> and \<open>k2 \<in> K\<close> show "a \<in> h ` derived_set G K" by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   428
    qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   429
  next
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   430
    show "h ` (derived_set G K) \<subseteq> derived_set H (h ` K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   431
    proof
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   432
      fix a assume "a \<in> h ` (derived_set G K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   433
      then obtain k1 k2 where "k1 \<in> K" "k2 \<in> K" "a = h (k1 \<otimes> k2 \<otimes> inv k1 \<otimes> inv k2)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   434
        by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   435
      hence "a = (h k1) \<otimes>\<^bsub>H\<^esub> (h k2) \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k1) \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k2)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   436
        using assms by (simp add: subset_iff)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   437
      from this \<open>k1 \<in> K\<close> and \<open>k2 \<in> K\<close> show "a \<in> derived_set H (h ` K)" by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   438
    qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   439
  qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   440
  thus ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   441
    unfolding derived_def using generate_img[OF G.derived_set_in_carrier[OF assms]] by simp
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   444
lemma (in group_hom) exp_of_derived_img:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   445
  assumes "K \<subseteq> carrier G" shows "(derived H ^^ n) (h ` K) = h ` ((derived G ^^ n) K)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   446
  using derived_img[OF G.exp_of_derived_in_carrier[OF assms]] by (induct n) (auto)
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   448
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   449
subsubsection \<open>Generated subgroup of a group\<close>
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   450
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   451
definition subgroup_generated :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> ('a, 'b) monoid_scheme"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   452
  where "subgroup_generated G S = G\<lparr>carrier := generate G (carrier G \<inter> S)\<rparr>"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   453
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   454
lemma carrier_subgroup_generated: "carrier (subgroup_generated G S) = generate G (carrier G \<inter> S)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   455
  by (auto simp: subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   456
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   457
lemma (in group) subgroup_generated_subset_carrier_subset:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   458
   "S \<subseteq> carrier G \<Longrightarrow> S \<subseteq> carrier(subgroup_generated G S)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   459
  by (simp add: Int_absorb1 carrier_subgroup_generated generate.incl subsetI)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   460
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   461
lemma (in group) subgroup_generated_minimal:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   462
   "\<lbrakk>subgroup H G; S \<subseteq> H\<rbrakk> \<Longrightarrow> carrier(subgroup_generated G S) \<subseteq> H"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   463
  by (simp add: carrier_subgroup_generated generate_subgroup_incl le_infI2)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   464
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   465
lemma (in group) carrier_subgroup_generated_subset:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   466
   "carrier (subgroup_generated G A) \<subseteq> carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   467
  apply (clarsimp simp: carrier_subgroup_generated)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   468
  by (meson Int_lower1 generate_in_carrier)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   469
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   470
lemma (in group) group_subgroup_generated [simp]: "group (subgroup_generated G S)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   471
    unfolding subgroup_generated_def
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   472
    by (simp add: generate_is_subgroup subgroup_imp_group)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   473
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   474
lemma (in group) abelian_subgroup_generated:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   475
  assumes "comm_group G"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   476
  shows "comm_group (subgroup_generated G S)" (is "comm_group ?GS")
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   477
proof (rule group.group_comm_groupI)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   478
  show "Group.group ?GS"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   479
    by simp
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   480
next
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   481
  fix x y
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   482
  assume "x \<in> carrier ?GS" "y \<in> carrier ?GS"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   483
  with assms show "x \<otimes>\<^bsub>?GS\<^esub> y = y \<otimes>\<^bsub>?GS\<^esub> x"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   484
    apply (simp add: subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   485
    by (meson Int_lower1 comm_groupE(4) generate_in_carrier)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   486
qed
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   487
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   488
lemma (in group) subgroup_of_subgroup_generated:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   489
  assumes "H \<subseteq> B" "subgroup H G"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   490
    shows "subgroup H (subgroup_generated G B)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   491
proof unfold_locales
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   492
  fix x
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   493
  assume "x \<in> H"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   494
  with assms show "inv\<^bsub>subgroup_generated G B\<^esub> x \<in> H"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   495
    unfolding subgroup_generated_def
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   496
    by (metis IntI Int_commute Int_lower2 generate.incl generate_is_subgroup m_inv_consistent subgroup_def subsetCE)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   497
next
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   498
  show "H \<subseteq> carrier (subgroup_generated G B)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   499
    using assms apply (auto simp: carrier_subgroup_generated)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   500
    by (metis Int_iff generate.incl inf.orderE subgroup.mem_carrier)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   501
qed (use assms in \<open>auto simp: subgroup_generated_def subgroup_def\<close>)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   502
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   503
lemma carrier_subgroup_generated_alt:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   504
  assumes "Group.group G" "S \<subseteq> carrier G"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   505
  shows "carrier (subgroup_generated G S) = \<Inter>{H. subgroup H G \<and> carrier G \<inter> S \<subseteq> H}"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   506
  using assms by (auto simp: group.generate_minimal subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   507
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   508
lemma one_subgroup_generated [simp]: "\<one>\<^bsub>subgroup_generated G S\<^esub> = \<one>\<^bsub>G\<^esub>"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   509
  by (auto simp: subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   510
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   511
lemma mult_subgroup_generated [simp]: "mult (subgroup_generated G S) = mult G"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   512
  by (auto simp: subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   513
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   514
lemma (in group) inv_subgroup_generated [simp]:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   515
  assumes "f \<in> carrier (subgroup_generated G S)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   516
  shows "inv\<^bsub>subgroup_generated G S\<^esub> f = inv f"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   517
proof (rule group.inv_equality)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   518
  show "Group.group (subgroup_generated G S)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   519
    by simp
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   520
  have [simp]: "f \<in> carrier G"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   521
    by (metis Int_lower1 assms carrier_subgroup_generated generate_in_carrier)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   522
  show "inv f \<otimes>\<^bsub>subgroup_generated G S\<^esub> f = \<one>\<^bsub>subgroup_generated G S\<^esub>"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   523
    by (simp add: assms)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   524
  show "f \<in> carrier (subgroup_generated G S)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   525
    using assms by (simp add: generate.incl subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   526
  show "inv f \<in> carrier (subgroup_generated G S)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   527
    using assms by (simp add: subgroup_generated_def generate_m_inv_closed)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   528
qed
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   529
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   530
lemma subgroup_generated_restrict [simp]:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   531
   "subgroup_generated G (carrier G \<inter> S) = subgroup_generated G S"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   532
  by (simp add: subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   533
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   534
lemma (in subgroup) carrier_subgroup_generated_subgroup [simp]:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   535
  "carrier (subgroup_generated G H) = H"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   536
  by (auto simp: generate.incl carrier_subgroup_generated elim: generate.induct)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   537
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   538
lemma (in group) subgroup_subgroup_generated_iff:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   539
   "subgroup H (subgroup_generated G B) \<longleftrightarrow> subgroup H G \<and> H \<subseteq> carrier(subgroup_generated G B)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   540
 (is "?lhs = ?rhs")
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   541
proof
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   542
  assume L: ?lhs
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   543
  then have Hsub: "H \<subseteq> generate G (carrier G \<inter> B)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   544
    by (simp add: subgroup_def subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   545
  then have H: "H \<subseteq> carrier G" "H \<subseteq> carrier(subgroup_generated G B)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   546
    unfolding carrier_subgroup_generated
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   547
    using generate_incl by blast+
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   548
  with Hsub have "subgroup H G"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   549
    by (metis Int_commute Int_lower2 L carrier_subgroup_generated generate_consistent
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   550
        generate_is_subgroup inf.orderE subgroup.carrier_subgroup_generated_subgroup subgroup_generated_def)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   551
  with H show ?rhs
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   552
    by blast
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   553
next
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   554
  assume ?rhs
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   555
  then show ?lhs
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   556
    by (simp add: generate_is_subgroup subgroup_generated_def subgroup_incl)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   557
qed
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   558
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   559
lemma (in group) subgroup_subgroup_generated:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   560
   "subgroup (carrier(subgroup_generated G S)) G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   561
  using group.subgroup_self group_subgroup_generated subgroup_subgroup_generated_iff by blast
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   562
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   563
lemma pow_subgroup_generated:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   564
   "pow (subgroup_generated G S) = (pow G :: 'a \<Rightarrow> nat \<Rightarrow> 'a)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   565
proof -
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   566
  have "x [^]\<^bsub>subgroup_generated G S\<^esub> n = x [^]\<^bsub>G\<^esub> n" for x and n::nat
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   567
    by (induction n) auto
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   568
  then show ?thesis
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   569
    by force
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   570
qed
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   571
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   572
lemma (in group) subgroup_generated2 [simp]: "subgroup_generated (subgroup_generated G S) S = subgroup_generated G S"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   573
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   574
  have *: "\<And>A. carrier G \<inter> A \<subseteq> carrier (subgroup_generated (subgroup_generated G A) A)"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 70039
diff changeset
   575
    by (metis (no_types, opaque_lifting) Int_assoc carrier_subgroup_generated generate.incl inf.order_iff subset_iff)
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   576
  show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   577
  apply (auto intro!: monoid.equality)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   578
    using group.carrier_subgroup_generated_subset group_subgroup_generated apply blast
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 70039
diff changeset
   579
     apply (metis (no_types, opaque_lifting) "*" group.subgroup_subgroup_generated group_subgroup_generated subgroup_generated_minimal
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   580
        subgroup_generated_restrict subgroup_subgroup_generated_iff subset_eq)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   581
    apply (simp add: subgroup_generated_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   582
    done
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   583
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   584
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   585
lemma (in group) int_pow_subgroup_generated:
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   586
  fixes n::int
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   587
  assumes "x \<in> carrier (subgroup_generated G S)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   588
  shows "x [^]\<^bsub>subgroup_generated G S\<^esub> n = x [^]\<^bsub>G\<^esub> n"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   589
proof -
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   590
  have "x [^] nat (- n) \<in> carrier (subgroup_generated G S)"
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   591
    by (metis assms group.is_monoid group_subgroup_generated monoid.nat_pow_closed pow_subgroup_generated)
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   592
  then show ?thesis
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   593
    by (metis group.inv_subgroup_generated int_pow_def2 is_group pow_subgroup_generated)
69749
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   594
qed
10e48c47a549 some new results in group theory
paulson <lp15@cam.ac.uk>
parents: 69122
diff changeset
   595
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   596
lemma kernel_from_subgroup_generated [simp]:
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   597
  "subgroup S G \<Longrightarrow> kernel (subgroup_generated G S) H f = kernel G H f \<inter> S"
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   598
  using subgroup.carrier_subgroup_generated_subgroup subgroup.subset
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   599
  by (fastforce simp add: kernel_def set_eq_iff)
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   600
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   601
lemma kernel_to_subgroup_generated [simp]:
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   602
  "kernel G (subgroup_generated H S) f = kernel G H f"
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   603
  by (simp add: kernel_def)
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69749
diff changeset
   604
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   605
subsection \<open>And homomorphisms\<close>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   606
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   607
lemma (in group) hom_from_subgroup_generated:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   608
   "h \<in> hom G H \<Longrightarrow> h \<in> hom(subgroup_generated G A) H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   609
  apply (simp add: hom_def carrier_subgroup_generated Pi_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   610
  apply (metis group.generate_in_carrier inf_le1 is_group)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   611
  done
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   612
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   613
lemma hom_into_subgroup:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   614
   "\<lbrakk>h \<in> hom G G'; h ` (carrier G) \<subseteq> H\<rbrakk> \<Longrightarrow> h \<in> hom G (subgroup_generated G' H)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   615
  by (auto simp: hom_def carrier_subgroup_generated Pi_iff generate.incl image_subset_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   616
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   617
lemma hom_into_subgroup_eq_gen:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   618
    "group G \<Longrightarrow>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   619
     h \<in> hom K (subgroup_generated G H)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   620
 \<longleftrightarrow> h \<in> hom K G \<and> h ` (carrier K) \<subseteq> carrier(subgroup_generated G H)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   621
  using group.carrier_subgroup_generated_subset [of G H] by (auto simp: hom_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   622
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   623
lemma hom_into_subgroup_eq:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   624
   "\<lbrakk>subgroup H G; group G\<rbrakk>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   625
    \<Longrightarrow> (h \<in> hom K (subgroup_generated G H) \<longleftrightarrow> h \<in> hom K G \<and> h ` (carrier K) \<subseteq> H)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   626
  by (simp add: hom_into_subgroup_eq_gen image_subset_iff subgroup.carrier_subgroup_generated_subgroup)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   627
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   628
lemma (in group_hom) hom_between_subgroups:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   629
  assumes "h ` A \<subseteq> B"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   630
  shows "h \<in> hom (subgroup_generated G A) (subgroup_generated H B)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   631
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   632
  have [simp]: "group G" "group H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   633
    by (simp_all add: G.is_group H.is_group)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   634
  have "x \<in> generate G (carrier G \<inter> A) \<Longrightarrow> h x \<in> generate H (carrier H \<inter> B)" for x
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   635
  proof (induction x rule: generate.induct)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   636
    case (incl h)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   637
    then show ?case
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   638
      by (meson IntE IntI assms generate.incl hom_closed image_subset_iff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   639
  next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   640
    case (inv h)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   641
    then show ?case
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   642
      by (metis G.inv_closed G.inv_inv IntE IntI assms generate.simps hom_inv image_subset_iff local.inv_closed)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   643
  next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   644
    case (eng h1 h2)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   645
    then show ?case
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   646
      by (metis G.generate_in_carrier generate.simps inf.cobounded1 local.hom_mult)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   647
  qed (auto simp: generate.intros)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   648
  then have "h ` carrier (subgroup_generated G A) \<subseteq> carrier (subgroup_generated H B)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   649
    using group.carrier_subgroup_generated_subset [of G A]
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   650
    by (auto simp: carrier_subgroup_generated)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   651
  then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   652
    by (simp add: hom_into_subgroup_eq_gen group.hom_from_subgroup_generated homh)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   653
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   654
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   655
lemma (in group_hom) subgroup_generated_by_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   656
  assumes "S \<subseteq> carrier G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   657
  shows "carrier (subgroup_generated H (h ` S)) = h ` (carrier(subgroup_generated G S))"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   658
proof
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   659
  have "x \<in> generate H (carrier H \<inter> h ` S) \<Longrightarrow> x \<in> h ` generate G (carrier G \<inter> S)" for x
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   660
  proof (erule generate.induct)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   661
    show "\<one>\<^bsub>H\<^esub> \<in> h ` generate G (carrier G \<inter> S)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   662
      using generate.one by force
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   663
  next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   664
    fix f
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   665
    assume "f \<in> carrier H \<inter> h ` S"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   666
    with assms show "f \<in> h ` generate G (carrier G \<inter> S)" "inv\<^bsub>H\<^esub> f \<in> h ` generate G (carrier G \<inter> S)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   667
       apply (auto simp: Int_absorb1 generate.incl)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   668
      apply (metis generate.simps hom_inv imageI subsetCE)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   669
      done
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   670
  next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   671
    fix h1 h2
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   672
    assume
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   673
      "h1 \<in> generate H (carrier H \<inter> h ` S)" "h1 \<in> h ` generate G (carrier G \<inter> S)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   674
      "h2 \<in> generate H (carrier H \<inter> h ` S)" "h2 \<in> h ` generate G (carrier G \<inter> S)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   675
    then show "h1 \<otimes>\<^bsub>H\<^esub> h2 \<in> h ` generate G (carrier G \<inter> S)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   676
      using H.subgroupE(4) group.generate_is_subgroup subgroup_img_is_subgroup
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   677
      by (simp add: G.generate_is_subgroup)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   678
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   679
  then
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   680
  show "carrier (subgroup_generated H (h ` S)) \<subseteq> h ` carrier (subgroup_generated G S)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   681
    by (auto simp: carrier_subgroup_generated)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   682
next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   683
  have "h ` S \<subseteq> carrier H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   684
    by (metis (no_types) assms hom_closed image_subset_iff subsetCE)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   685
  then show "h ` carrier (subgroup_generated G S) \<subseteq> carrier (subgroup_generated H (h ` S))"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   686
    apply (clarsimp simp: carrier_subgroup_generated)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   687
    by (metis Int_absorb1 assms generate_img imageI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   688
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   689
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   690
lemma (in group_hom) iso_between_subgroups:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   691
  assumes "h \<in> iso G H" "S \<subseteq> carrier G" "h ` S = T"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   692
  shows "h \<in> iso (subgroup_generated G S) (subgroup_generated H T)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   693
  using assms
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   694
  by (metis G.carrier_subgroup_generated_subset Group.iso_iff hom_between_subgroups inj_on_subset subgroup_generated_by_image subsetI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   695
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   696
lemma (in group) subgroup_generated_group_carrier:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   697
  "subgroup_generated G (carrier G) = G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   698
proof (rule monoid.equality)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   699
  show "carrier (subgroup_generated G (carrier G)) = carrier G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   700
    by (simp add: subgroup.carrier_subgroup_generated_subgroup subgroup_self)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   701
qed (auto simp: subgroup_generated_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   702
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   703
lemma iso_onto_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   704
  assumes "group G" "group H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   705
  shows
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   706
   "f \<in> iso G (subgroup_generated H (f ` carrier G)) \<longleftrightarrow> f \<in> hom G H \<and> inj_on f (carrier G)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   707
  using assms
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   708
  apply (auto simp: iso_def bij_betw_def hom_into_subgroup_eq_gen carrier_subgroup_generated hom_carrier generate.incl Int_absorb1 Int_absorb2)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   709
  by (metis group.generateI group.subgroupE(1) group.subgroup_self group_hom.generate_img group_hom.intro group_hom_axioms.intro)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   710
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   711
lemma (in group) iso_onto_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   712
   "group H \<Longrightarrow> f \<in> iso G (subgroup_generated H (f ` carrier G)) \<longleftrightarrow> f \<in> mon G H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   713
  by (simp add: mon_def epi_def hom_into_subgroup_eq_gen iso_onto_image)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   714
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   715
end