src/HOL/Algebra/Elementary_Groups.thy
author wenzelm
Tue, 05 Nov 2019 14:28:00 +0100
changeset 71047 87c132cf5860
parent 70039 733e256ecdf3
child 72630 4167d3d3d478
permissions -rw-r--r--
more options;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
section \<open>Elementary Group Constructions\<close>
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
(*  Title:      HOL/Algebra/Elementary_Groups.thy
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
    Author:     LC Paulson, ported from HOL Light
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
*)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory Elementary_Groups
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
     8
imports Generated_Groups Multiplicative_Group "HOL-Library.Infinite_Set"
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
begin
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
subsection\<open>Direct sum/product lemmas\<close>
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
locale group_disjoint_sum = group G + AG: subgroup A G + BG: subgroup B G for G (structure) and A B
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
begin
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
lemma subset_one: "A \<inter> B \<subseteq> {\<one>} \<longleftrightarrow> A \<inter> B = {\<one>}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  by  auto
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
lemma sub_id_iff: "A \<inter> B \<subseteq> {\<one>} \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = \<one> \<longrightarrow> x = \<one> \<and> y = \<one>)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
        (is "?lhs = ?rhs")
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
  have "?lhs = (\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> inv y = \<one> \<longrightarrow> x = \<one> \<and> inv y = \<one>)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
  proof (intro ballI iffI impI)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
    fix x y
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
    assume "A \<inter> B \<subseteq> {\<one>}" "x \<in> A" "y \<in> B" "x \<otimes> inv y = \<one>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
    then have "y = x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
      using group.inv_equality group_l_invI by fastforce
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
    then show "x = \<one> \<and> inv y = \<one>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
      using \<open>A \<inter> B \<subseteq> {\<one>}\<close> \<open>x \<in> A\<close> \<open>y \<in> B\<close> by fastforce
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  next
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
    assume "\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> inv y = \<one> \<longrightarrow> x = \<one> \<and> inv y = \<one>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
    then show "A \<inter> B \<subseteq> {\<one>}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
      by auto
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
  qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
  also have "\<dots> = ?rhs"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
    by (metis BG.mem_carrier BG.subgroup_axioms inv_inv subgroup_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
  finally show ?thesis .
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
lemma cancel: "A \<inter> B \<subseteq> {\<one>} \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. \<forall>x'\<in>A. \<forall>y'\<in>B. x \<otimes> y = x' \<otimes> y' \<longrightarrow> x = x' \<and> y = y')"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
        (is "?lhs = ?rhs")
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
  have "(\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = \<one> \<longrightarrow> x = \<one> \<and> y = \<one>) = ?rhs"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
    (is "?med = _")
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
  proof (intro ballI iffI impI)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
    fix x y x' y'
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
    assume * [rule_format]: "\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = \<one> \<longrightarrow> x = \<one> \<and> y = \<one>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
      and AB: "x \<in> A" "y \<in> B" "x' \<in> A" "y' \<in> B" and eq: "x \<otimes> y = x' \<otimes> y'"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
    then have carr: "x \<in> carrier G" "x' \<in> carrier G" "y \<in> carrier G" "y' \<in> carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
      using AG.subset BG.subset by auto
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
    then have "inv x' \<otimes> x \<otimes> (y \<otimes> inv y') = inv x' \<otimes> (x \<otimes> y) \<otimes> inv y'"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
      by (simp add: m_assoc)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
    also have "\<dots> = \<one>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
      using carr  by (simp add: eq) (simp add: m_assoc)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
    finally have 1: "inv x' \<otimes> x \<otimes> (y \<otimes> inv y') = \<one>" .
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
    show "x = x' \<and> y = y'"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
      using * [OF _ _ 1] AB by simp (metis carr inv_closed inv_inv local.inv_equality)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
  next
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
    fix x y
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
    assume * [rule_format]: "\<forall>x\<in>A. \<forall>y\<in>B. \<forall>x'\<in>A. \<forall>y'\<in>B. x \<otimes> y = x' \<otimes> y' \<longrightarrow> x = x' \<and> y = y'"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
      and xy: "x \<in> A" "y \<in> B" "x \<otimes> y = \<one>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
    show "x = \<one> \<and> y = \<one>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
      by (rule *) (use xy in auto)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
  qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
  then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
    by (simp add: sub_id_iff)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
lemma commuting_imp_normal1:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
  assumes sub: "carrier G \<subseteq> A <#> B"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
     and mult: "\<And>x y. \<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> x \<otimes> y = y \<otimes> x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
   shows "A \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
  have AB: "A \<subseteq> carrier G \<and> B \<subseteq> carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
    by (simp add: AG.subset BG.subset)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
  have "A #> x = x <# A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
    if x: "x \<in> carrier G" for x
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
  proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
    obtain a b where xeq: "x = a \<otimes> b" and "a \<in> A" "b \<in> B" and carr: "a \<in> carrier G" "b \<in> carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
      using x sub AB by (force simp: set_mult_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
    have Ab: "A <#> {b} = {b} <#> A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
      using AB \<open>a \<in> A\<close> \<open>b \<in> B\<close> mult
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
      by (force simp: set_mult_def m_assoc subset_iff)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
    have "A #> x = A <#> {a \<otimes> b}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
      by (auto simp: l_coset_eq_set_mult r_coset_eq_set_mult xeq)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
    also have "\<dots> = A <#> {a} <#> {b}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
      using AB \<open>a \<in> A\<close> \<open>b \<in> B\<close>
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
      by (auto simp: set_mult_def m_assoc subset_iff)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
    also have "\<dots> = {a} <#> A <#> {b}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
      by (metis AG.rcos_const AG.subgroup_axioms \<open>a \<in> A\<close> coset_join3 is_group l_coset_eq_set_mult r_coset_eq_set_mult subgroup.mem_carrier)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
    also have "\<dots> = {a} <#> {b} <#> A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
      by (simp add: is_group carr group.set_mult_assoc AB Ab)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
    also have "\<dots> = {x} <#> A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
      by (auto simp: set_mult_def xeq)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
    finally show "A #> x = x <# A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
      by (simp add: l_coset_eq_set_mult)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
  qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
  then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
    by (auto simp: normal_def normal_axioms_def AG.subgroup_axioms is_group)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
lemma commuting_imp_normal2:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
  assumes"carrier G \<subseteq> A <#> B"  "\<And>x y. \<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> x \<otimes> y = y \<otimes> x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  shows "B \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
proof (rule group_disjoint_sum.commuting_imp_normal1)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
  show "group_disjoint_sum G B A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
  proof qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
next
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
  show "carrier G \<subseteq> B <#> A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
    using BG.subgroup_axioms assms commut_normal commuting_imp_normal1 by blast
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
qed (use assms in auto)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
lemma (in group) normal_imp_commuting:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
  assumes "A \<lhd> G" "B \<lhd> G" "A \<inter> B \<subseteq> {\<one>}" "x \<in> A" "y \<in> B"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
  shows "x \<otimes> y = y \<otimes> x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
  interpret AG: normal A G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
    using assms by auto
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  interpret BG: normal B G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
    using assms by auto
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
  interpret group_disjoint_sum G A B
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
  proof qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
  have * [rule_format]: "(\<forall>x\<in>A. \<forall>y\<in>B. \<forall>x'\<in>A. \<forall>y'\<in>B. x \<otimes> y = x' \<otimes> y' \<longrightarrow> x = x' \<and> y = y')"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
    using cancel assms by (auto simp: normal_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  have carr: "x \<in> carrier G" "y \<in> carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
    using assms AG.subset BG.subset by auto
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
    using * [of x _ _ y] AG.coset_eq [rule_format, of y] BG.coset_eq [rule_format, of x]
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
    by (clarsimp simp: l_coset_def r_coset_def set_eq_iff) (metis \<open>x \<in> A\<close> \<open>y \<in> B\<close>)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
lemma normal_eq_commuting:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
  assumes "carrier G \<subseteq> A <#> B" "A \<inter> B \<subseteq> {\<one>}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
  shows "A \<lhd> G \<and> B \<lhd> G \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = y \<otimes> x)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
  by (metis assms commuting_imp_normal1 commuting_imp_normal2 normal_imp_commuting)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
lemma (in group) hom_group_mul_rev:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
  assumes "(\<lambda>(x,y). x \<otimes> y) \<in> hom (subgroup_generated G A \<times>\<times> subgroup_generated G B) G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
          (is "?h \<in> hom ?P G")
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
   and "x \<in> carrier G" "y \<in> carrier G" "x \<in> A" "y \<in> B"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
 shows "x \<otimes> y = y \<otimes> x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
  interpret P: group_hom ?P G ?h
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
    by (simp add: assms DirProd_group group_hom.intro group_hom_axioms.intro is_group)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
  have xy: "(x,y) \<in> carrier ?P"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
    by (auto simp: assms carrier_subgroup_generated generate.incl)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
  have "x \<otimes> (x \<otimes> (y \<otimes> y)) = x \<otimes> (y \<otimes> (x \<otimes> y))"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
    using P.hom_mult [OF xy xy] by (simp add: m_assoc assms)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
  then have "x \<otimes> (y \<otimes> y) = y \<otimes> (x \<otimes> y)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
    using assms by simp
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
    by (simp add: assms flip: m_assoc)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
lemma hom_group_mul_eq:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
   "(\<lambda>(x,y). x \<otimes> y) \<in> hom (subgroup_generated G A \<times>\<times> subgroup_generated G B) G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
     \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = y \<otimes> x)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
         (is "?lhs = ?rhs")
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
proof
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
  assume ?lhs then show ?rhs
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
    using hom_group_mul_rev AG.subset BG.subset by blast
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
next
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
  assume R: ?rhs
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
  have subG: "generate G (carrier G \<inter> A) \<subseteq> carrier G" for A
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
    by (simp add: generate_incl)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
  have *: "x \<otimes> u \<otimes> (y \<otimes> v) = x \<otimes> y \<otimes> (u \<otimes> v)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
    if eq [rule_format]: "\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = y \<otimes> x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
      and gen: "x \<in> generate G (carrier G \<inter> A)" "y \<in> generate G (carrier G \<inter> B)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
      "u \<in> generate G (carrier G \<inter> A)" "v \<in> generate G (carrier G \<inter> B)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
    for x y u v
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
  proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
    have "u \<otimes> y = y \<otimes> u"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
      by (metis AG.carrier_subgroup_generated_subgroup BG.carrier_subgroup_generated_subgroup carrier_subgroup_generated eq that(3) that(4))
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
    then have "x \<otimes> u \<otimes> y = x \<otimes> y \<otimes> u"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
      using gen by (simp add: m_assoc subsetD [OF subG])
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
    then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
      using gen by (simp add: subsetD [OF subG] flip: m_assoc)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
  show ?lhs
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
    using R by (auto simp: hom_def carrier_subgroup_generated subsetD [OF subG] *)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
lemma epi_group_mul_eq:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
   "(\<lambda>(x,y). x \<otimes> y) \<in> epi (subgroup_generated G A \<times>\<times> subgroup_generated G B) G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
     \<longleftrightarrow> A <#> B = carrier G \<and> (\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = y \<otimes> x)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
  have subGA: "generate G (carrier G \<inter> A) \<subseteq> A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
    by (simp add: AG.subgroup_axioms generate_subgroup_incl)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
  have subGB: "generate G (carrier G \<inter> B) \<subseteq> B"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
    by (simp add: BG.subgroup_axioms generate_subgroup_incl)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
  have "(((\<lambda>(x, y). x \<otimes> y) ` (generate G (carrier G \<inter> A) \<times> generate G (carrier G \<inter> B)))) = ((A <#> B))"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
    by (auto simp: set_mult_def generate.incl pair_imageI dest: subsetD [OF subGA] subsetD [OF subGB])
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
  then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
    by (auto simp: epi_def hom_group_mul_eq carrier_subgroup_generated)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
lemma mon_group_mul_eq:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
    "(\<lambda>(x,y). x \<otimes> y) \<in> mon (subgroup_generated G A \<times>\<times> subgroup_generated G B) G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
     \<longleftrightarrow> A \<inter> B = {\<one>} \<and> (\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = y \<otimes> x)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
  have subGA: "generate G (carrier G \<inter> A) \<subseteq> A"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
    by (simp add: AG.subgroup_axioms generate_subgroup_incl)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
  have subGB: "generate G (carrier G \<inter> B) \<subseteq> B"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
    by (simp add: BG.subgroup_axioms generate_subgroup_incl)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
  show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
    apply (auto simp: mon_def hom_group_mul_eq simp flip: subset_one)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
     apply (simp_all (no_asm_use) add: inj_on_def AG.carrier_subgroup_generated_subgroup BG.carrier_subgroup_generated_subgroup)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
    using cancel apply blast+
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
    done
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
lemma iso_group_mul_alt:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
    "(\<lambda>(x,y). x \<otimes> y) \<in> iso (subgroup_generated G A \<times>\<times> subgroup_generated G B) G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
     \<longleftrightarrow> A \<inter> B = {\<one>} \<and> A <#> B = carrier G \<and> (\<forall>x\<in>A. \<forall>y\<in>B. x \<otimes> y = y \<otimes> x)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
  by (auto simp: iso_iff_mon_epi mon_group_mul_eq epi_group_mul_eq)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
lemma iso_group_mul_eq:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
    "(\<lambda>(x,y). x \<otimes> y) \<in> iso (subgroup_generated G A \<times>\<times> subgroup_generated G B) G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
     \<longleftrightarrow> A \<inter> B = {\<one>} \<and> A <#> B = carrier G \<and> A \<lhd> G \<and> B \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
  by (simp add: iso_group_mul_alt normal_eq_commuting cong: conj_cong)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
lemma (in group) iso_group_mul_gen:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
  assumes "A \<lhd> G" "B \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
  shows "(\<lambda>(x,y). x \<otimes> y) \<in> iso (subgroup_generated G A \<times>\<times> subgroup_generated G B) G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
     \<longleftrightarrow> A \<inter> B \<subseteq> {\<one>} \<and> A <#> B = carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
  interpret group_disjoint_sum G A B
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
    using assms by (auto simp: group_disjoint_sum_def normal_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
  show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
    by (simp add: subset_one iso_group_mul_eq assms)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
lemma iso_group_mul:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
  assumes "comm_group G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
  shows "((\<lambda>(x,y). x \<otimes> y) \<in> iso (DirProd (subgroup_generated G A) (subgroup_generated G B)) G
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
     \<longleftrightarrow> A \<inter> B \<subseteq> {\<one>} \<and> A <#> B = carrier G)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
proof (rule iso_group_mul_gen)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
  interpret comm_group
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
    by (rule assms)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
  show "A \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
    by (simp add: AG.subgroup_axioms subgroup_imp_normal)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
  show "B \<lhd> G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
    by (simp add: BG.subgroup_axioms subgroup_imp_normal)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
end
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
subsection\<open>The one-element group on a given object\<close>
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
definition singleton_group :: "'a \<Rightarrow> 'a monoid"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
  where "singleton_group a = \<lparr>carrier = {a}, monoid.mult = (\<lambda>x y. a), one = a\<rparr>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
lemma singleton_group [simp]: "group (singleton_group a)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
  unfolding singleton_group_def by (auto intro: groupI)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
lemma singleton_abelian_group [simp]: "comm_group (singleton_group a)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
  by (metis group.group_comm_groupI monoid.simps(1) singleton_group singleton_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
lemma carrier_singleton_group [simp]: "carrier (singleton_group a) = {a}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
  by (auto simp: singleton_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
lemma (in group) hom_into_singleton_iff [simp]:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  "h \<in> hom G (singleton_group a) \<longleftrightarrow> h \<in> carrier G \<rightarrow> {a}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
  by (auto simp: hom_def singleton_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
declare group.hom_into_singleton_iff [simp]
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
lemma (in group) id_hom_singleton: "id \<in> hom (singleton_group \<one>) G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
  by (simp add: hom_def singleton_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
subsection\<open>Similarly, trivial groups\<close>
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
definition trivial_group :: "('a, 'b) monoid_scheme \<Rightarrow> bool"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
  where "trivial_group G \<equiv> group G \<and> carrier G = {one G}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
lemma trivial_imp_finite_group:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
   "trivial_group G \<Longrightarrow> finite(carrier G)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
  by (simp add: trivial_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
lemma trivial_singleton_group [simp]: "trivial_group(singleton_group a)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
  by (metis monoid.simps(2) partial_object.simps(1) singleton_group singleton_group_def trivial_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
lemma (in group) trivial_group_subset:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
   "trivial_group G \<longleftrightarrow> carrier G \<subseteq> {one G}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
  using is_group trivial_group_def by fastforce
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
lemma (in group) trivial_group: "trivial_group G \<longleftrightarrow> (\<exists>a. carrier G = {a})"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
  unfolding trivial_group_def using one_closed is_group by fastforce
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
lemma (in group) trivial_group_alt:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
   "trivial_group G \<longleftrightarrow> (\<exists>a. carrier G \<subseteq> {a})"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
  by (auto simp: trivial_group)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
lemma (in group) trivial_group_subgroup_generated:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
  assumes "S \<subseteq> {one G}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
  shows "trivial_group(subgroup_generated G S)"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
proof -
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
  have "carrier (subgroup_generated G S) \<subseteq> {\<one>}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
    using generate_empty generate_one subset_singletonD assms
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
    by (fastforce simp add: carrier_subgroup_generated)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
  then show ?thesis
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
    by (simp add: group.trivial_group_subset)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
lemma (in group) trivial_group_subgroup_generated_eq:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
  "trivial_group(subgroup_generated G s) \<longleftrightarrow> carrier G \<inter> s \<subseteq> {one G}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
  apply (rule iffI)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
   apply (force simp: trivial_group_def carrier_subgroup_generated generate.incl)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
  by (metis subgroup_generated_restrict trivial_group_subgroup_generated)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
lemma isomorphic_group_triviality1:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
  assumes "G \<cong> H" "group H" "trivial_group G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
  shows "trivial_group H"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
  using assms
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
  by (auto simp: trivial_group_def is_iso_def iso_def group.is_monoid Group.group_def bij_betw_def hom_one)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
lemma isomorphic_group_triviality:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
  assumes "G \<cong> H" "group G" "group H"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
  shows "trivial_group G \<longleftrightarrow> trivial_group H"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
  by (meson assms group.iso_sym isomorphic_group_triviality1)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
lemma (in group_hom) kernel_from_trivial_group:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
   "trivial_group G \<Longrightarrow> kernel G H h = carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
  by (auto simp: trivial_group_def kernel_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
lemma (in group_hom) image_from_trivial_group:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
   "trivial_group G \<Longrightarrow> h ` carrier G = {one H}"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
  by (auto simp: trivial_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
lemma (in group_hom) kernel_to_trivial_group:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
   "trivial_group H \<Longrightarrow> kernel G H h = carrier G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
  unfolding kernel_def trivial_group_def
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
  using hom_closed by blast
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
subsection\<open>The additive group of integers\<close>
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
definition integer_group
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
  where "integer_group = \<lparr>carrier = UNIV, monoid.mult = (+), one = (0::int)\<rparr>"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
lemma group_integer_group [simp]: "group integer_group"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
  unfolding integer_group_def
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
proof (rule groupI; simp)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
  show "\<And>x::int. \<exists>y. y + x = 0"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
  by presburger
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
qed
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
lemma carrier_integer_group [simp]: "carrier integer_group = UNIV"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
  by (auto simp: integer_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
lemma one_integer_group [simp]: "\<one>\<^bsub>integer_group\<^esub> = 0"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
  by (auto simp: integer_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
lemma mult_integer_group [simp]: "x \<otimes>\<^bsub>integer_group\<^esub> y = x + y"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
  by (auto simp: integer_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
lemma inv_integer_group [simp]: "inv\<^bsub>integer_group\<^esub> x = -x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
  by (rule group.inv_equality [OF group_integer_group]) (auto simp: integer_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
lemma abelian_integer_group: "comm_group integer_group"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
  by (rule group.group_comm_groupI [OF group_integer_group]) (auto simp: integer_group_def)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
lemma group_nat_pow_integer_group [simp]:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
  fixes n::nat and x::int
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
  shows "pow integer_group x n = int n * x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
  by (induction n) (auto simp: integer_group_def algebra_simps)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
lemma group_int_pow_integer_group [simp]:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
  fixes n::int and x::int
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
  shows "pow integer_group x n = n * x"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
  by (simp add: int_pow_def2)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
lemma (in group) hom_integer_group_pow:
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
   "x \<in> carrier G \<Longrightarrow> pow G x \<in> hom integer_group G"
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
  by (rule homI) (auto simp: int_pow_mult)
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
70039
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   382
subsection\<open>Additive group of integers modulo n (n = 0 gives just the integers)\<close>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   383
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   384
definition integer_mod_group :: "nat \<Rightarrow> int monoid"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   385
  where
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   386
  "integer_mod_group n \<equiv>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   387
     if n = 0 then integer_group
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   388
     else \<lparr>carrier = {0..<int n}, monoid.mult = (\<lambda>x y. (x+y) mod int n), one = 0\<rparr>"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   389
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   390
lemma carrier_integer_mod_group:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   391
  "carrier(integer_mod_group n) = (if n=0 then UNIV else {0..<int n})"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   392
  by (simp add: integer_mod_group_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   393
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   394
lemma one_integer_mod_group[simp]: "one(integer_mod_group n) = 0"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   395
  by (simp add: integer_mod_group_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   396
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   397
lemma mult_integer_mod_group[simp]: "monoid.mult(integer_mod_group n) = (\<lambda>x y. (x + y) mod int n)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   398
  by (simp add: integer_mod_group_def integer_group_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   399
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   400
lemma group_integer_mod_group [simp]: "group (integer_mod_group n)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   401
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   402
  have *: "\<exists>y\<ge>0. y < int n \<and> (y + x) mod int n = 0" if "x < int n" "0 \<le> x" for x
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   403
  proof (cases "x=0")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   404
    case False
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   405
    with that show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   406
      by (rule_tac x="int n - x" in exI) auto
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   407
  qed (use that in auto)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   408
  show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   409
    apply (rule groupI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   410
        apply (auto simp: integer_mod_group_def Bex_def *, presburger+)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   411
    done
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   412
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   413
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   414
lemma inv_integer_mod_group[simp]:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   415
  "x \<in> carrier (integer_mod_group n) \<Longrightarrow> m_inv(integer_mod_group n) x = (-x) mod int n"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   416
  by (rule group.inv_equality [OF group_integer_mod_group]) (auto simp: integer_mod_group_def add.commute mod_add_right_eq)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   417
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   418
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   419
lemma pow_integer_mod_group [simp]:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   420
  fixes m::nat
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   421
  shows "pow (integer_mod_group n) x m = (int m * x) mod int n"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   422
proof (cases "n=0")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   423
  case False
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   424
  show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   425
    by (induction m) (auto simp: add.commute mod_add_right_eq distrib_left mult.commute)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   426
qed (simp add: integer_mod_group_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   427
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   428
lemma int_pow_integer_mod_group:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   429
  "pow (integer_mod_group n) x m = (m * x) mod int n"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   430
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   431
  have "inv\<^bsub>integer_mod_group n\<^esub> (- (m * x) mod int n) = m * x mod int n"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   432
    by (simp add: carrier_integer_mod_group mod_minus_eq)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   433
  then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   434
    by (simp add: int_pow_def2)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   435
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   436
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   437
lemma abelian_integer_mod_group [simp]: "comm_group(integer_mod_group n)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   438
  by (simp add: add.commute group.group_comm_groupI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   439
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   440
lemma integer_mod_group_0 [simp]: "0 \<in> carrier(integer_mod_group n)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   441
  by (simp add: integer_mod_group_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   442
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   443
lemma integer_mod_group_1 [simp]: "1 \<in> carrier(integer_mod_group n) \<longleftrightarrow> (n \<noteq> 1)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   444
  by (auto simp: integer_mod_group_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   445
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   446
lemma trivial_integer_mod_group: "trivial_group(integer_mod_group n) \<longleftrightarrow> n = 1"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   447
  (is "?lhs = ?rhs")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   448
proof
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   449
  assume ?lhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   450
  then show ?rhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   451
  by (simp add: trivial_group_def carrier_integer_mod_group set_eq_iff split: if_split_asm) (presburger+)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   452
next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   453
  assume ?rhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   454
  then show ?lhs
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   455
    by (force simp: trivial_group_def carrier_integer_mod_group)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   456
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   457
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   458
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   459
subsection\<open>Cyclic groups\<close>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   460
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   461
lemma (in group) subgroup_of_powers:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   462
  "x \<in> carrier G \<Longrightarrow> subgroup (range (\<lambda>n::int. x [^] n)) G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   463
  apply (auto simp: subgroup_def image_iff simp flip: int_pow_mult int_pow_neg)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   464
  apply (metis group.int_pow_diff int_pow_closed is_group r_inv)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   465
  done
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   466
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   467
lemma (in group) carrier_subgroup_generated_by_singleton:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   468
  assumes "x \<in> carrier G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   469
  shows "carrier(subgroup_generated G {x}) = (range (\<lambda>n::int. x [^] n))"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   470
proof
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   471
  show "carrier (subgroup_generated G {x}) \<subseteq> range (\<lambda>n::int. x [^] n)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   472
  proof (rule subgroup_generated_minimal)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   473
    show "subgroup (range (\<lambda>n::int. x [^] n)) G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   474
      using assms subgroup_of_powers by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   475
    show "{x} \<subseteq> range (\<lambda>n::int. x [^] n)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   476
      by clarify (metis assms int_pow_1 range_eqI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   477
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   478
  have x: "x \<in> carrier (subgroup_generated G {x})"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   479
    using assms subgroup_generated_subset_carrier_subset by auto
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   480
  show "range (\<lambda>n::int. x [^] n) \<subseteq> carrier (subgroup_generated G {x})"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   481
  proof clarify
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   482
    fix n :: "int"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   483
    show "x [^] n \<in> carrier (subgroup_generated G {x})"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   484
        by (simp add: x subgroup_int_pow_closed subgroup_subgroup_generated)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   485
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   486
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   487
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   488
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   489
definition cyclic_group
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   490
  where "cyclic_group G \<equiv> \<exists>x \<in> carrier G. subgroup_generated G {x} = G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   491
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   492
lemma (in group) cyclic_group:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   493
  "cyclic_group G \<longleftrightarrow> (\<exists>x \<in> carrier G. carrier G = range (\<lambda>n::int. x [^] n))"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   494
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   495
  have "\<And>x. \<lbrakk>x \<in> carrier G; carrier G = range (\<lambda>n::int. x [^] n)\<rbrakk>
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   496
         \<Longrightarrow> \<exists>x\<in>carrier G. subgroup_generated G {x} = G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   497
    by (rule_tac x=x in bexI) (auto simp: generate_pow subgroup_generated_def intro!: monoid.equality)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   498
  then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   499
    unfolding cyclic_group_def
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   500
    using carrier_subgroup_generated_by_singleton by fastforce
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   501
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   502
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   503
lemma cyclic_integer_group [simp]: "cyclic_group integer_group"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   504
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   505
  have *: "int n \<in> generate integer_group {1}" for n
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   506
  proof (induction n)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   507
    case 0
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   508
    then show ?case
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   509
    using generate.simps by force
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   510
  next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   511
    case (Suc n)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   512
    then show ?case
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   513
    by simp (metis generate.simps insert_subset integer_group_def monoid.simps(1) subsetI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   514
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   515
  have **: "i \<in> generate integer_group {1}" for i
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   516
  proof (cases i rule: int_cases)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   517
    case (nonneg n)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   518
    then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   519
      by (simp add: *)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   520
  next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   521
    case (neg n)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   522
    then have "-i \<in> generate integer_group {1}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   523
      by (metis "*" add.inverse_inverse)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   524
    then have "- (-i) \<in> generate integer_group {1}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   525
      by (metis UNIV_I group.generate_m_inv_closed group_integer_group integer_group_def inv_integer_group partial_object.select_convs(1) subsetI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   526
    then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   527
      by simp
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   528
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   529
  show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   530
    unfolding cyclic_group_def
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   531
    by (rule_tac x=1 in bexI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   532
       (auto simp: carrier_subgroup_generated ** intro: monoid.equality)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   533
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   534
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   535
lemma nontrivial_integer_group [simp]: "\<not> trivial_group integer_group"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   536
  using integer_mod_group_def trivial_integer_mod_group by presburger
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   537
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   538
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   539
lemma (in group) cyclic_imp_abelian_group:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   540
  "cyclic_group G \<Longrightarrow> comm_group G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   541
  apply (auto simp: cyclic_group comm_group_def is_group intro!: monoid_comm_monoidI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   542
  apply (metis add.commute int_pow_mult rangeI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   543
  done
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   544
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   545
lemma trivial_imp_cyclic_group:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   546
   "trivial_group G \<Longrightarrow> cyclic_group G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   547
  by (metis cyclic_group_def group.subgroup_generated_group_carrier insertI1 trivial_group_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   548
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   549
lemma (in group) cyclic_group_alt:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   550
   "cyclic_group G \<longleftrightarrow> (\<exists>x. subgroup_generated G {x} = G)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   551
proof safe
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   552
  fix x
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   553
  assume *: "subgroup_generated G {x} = G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   554
  show "cyclic_group G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   555
  proof (cases "x \<in> carrier G")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   556
    case True
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   557
    then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   558
      using \<open>subgroup_generated G {x} = G\<close> cyclic_group_def by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   559
  next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   560
    case False
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   561
    then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   562
      by (metis "*" Int_empty_right Int_insert_right_if0 carrier_subgroup_generated generate_empty trivial_group trivial_imp_cyclic_group)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   563
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   564
qed (auto simp: cyclic_group_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   565
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   566
lemma (in group) cyclic_group_generated:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   567
  "cyclic_group (subgroup_generated G {x})"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   568
  using group.cyclic_group_alt group_subgroup_generated subgroup_generated2 by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   569
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   570
lemma (in group) cyclic_group_epimorphic_image:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   571
  assumes "h \<in> epi G H" "cyclic_group G" "group H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   572
  shows "cyclic_group H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   573
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   574
  interpret h: group_hom
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   575
    using assms
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   576
    by (simp add: group_hom_def group_hom_axioms_def is_group epi_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   577
  obtain x where "x \<in> carrier G" and x: "carrier G = range (\<lambda>n::int. x [^] n)" and eq: "carrier H = h ` carrier G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   578
    using assms by (auto simp: cyclic_group epi_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   579
  have "h ` carrier G = range (\<lambda>n::int. h x [^]\<^bsub>H\<^esub> n)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   580
    by (metis (no_types, lifting) \<open>x \<in> carrier G\<close> h.hom_int_pow image_cong image_image x)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   581
  then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   582
    using \<open>x \<in> carrier G\<close> eq h.cyclic_group by blast
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   583
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   584
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   585
lemma isomorphic_group_cyclicity:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   586
   "\<lbrakk>G \<cong> H; group G; group H\<rbrakk> \<Longrightarrow> cyclic_group G \<longleftrightarrow> cyclic_group H"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   587
  by (meson ex_in_conv group.cyclic_group_epimorphic_image group.iso_sym is_iso_def iso_iff_mon_epi)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   588
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   589
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   590
lemma (in group)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   591
  assumes "x \<in> carrier G"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   592
  shows finite_cyclic_subgroup:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   593
        "finite(carrier(subgroup_generated G {x})) \<longleftrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> x [^] n = \<one>)" (is "?fin \<longleftrightarrow> ?nat1")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   594
    and infinite_cyclic_subgroup:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   595
        "infinite(carrier(subgroup_generated G {x})) \<longleftrightarrow> (\<forall>m n::nat. x [^] m = x [^] n \<longrightarrow> m = n)" (is "\<not> ?fin \<longleftrightarrow> ?nateq")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   596
    and finite_cyclic_subgroup_int:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   597
        "finite(carrier(subgroup_generated G {x})) \<longleftrightarrow> (\<exists>i::int. i \<noteq> 0 \<and> x [^] i = \<one>)" (is "?fin \<longleftrightarrow> ?int1")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   598
    and infinite_cyclic_subgroup_int:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   599
        "infinite(carrier(subgroup_generated G {x})) \<longleftrightarrow> (\<forall>i j::int. x [^] i = x [^] j \<longrightarrow> i = j)" (is "\<not> ?fin \<longleftrightarrow> ?inteq")
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   600
proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   601
  have 1: "\<not> ?fin" if ?nateq
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   602
  proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   603
    have "infinite (range (\<lambda>n::nat. x [^] n))"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   604
      using that range_inj_infinite [of "(\<lambda>n::nat. x [^] n)"] by (auto simp: inj_on_def)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   605
    moreover have "range (\<lambda>n::nat. x [^] n) \<subseteq> range (\<lambda>i::int. x [^] i)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   606
      apply clarify
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   607
      by (metis assms group.int_pow_neg int_pow_closed int_pow_neg_int is_group local.inv_equality nat_pow_closed r_inv rangeI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   608
    ultimately show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   609
      using carrier_subgroup_generated_by_singleton [OF assms] finite_subset by auto
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   610
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   611
  have 2: "m = n" if mn: "x [^] m = x [^] n" and eq [rule_format]: "?inteq" for m n::nat
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   612
    using eq [of "int m" "int n"]
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   613
    by (simp add: int_pow_int mn)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   614
  have 3: ?nat1 if non: "\<not> ?inteq"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   615
  proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   616
    obtain i j::int where eq: "x [^] i = x [^] j" and "i \<noteq> j"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   617
      using non by auto
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   618
    show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   619
    proof (cases i j rule: linorder_cases)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   620
      case less
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   621
      then have [simp]: "x [^] (j - i) = \<one>"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   622
        by (simp add: eq assms int_pow_diff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   623
      show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   624
        using less by (rule_tac x="nat (j-i)" in exI) auto
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   625
    next
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   626
      case greater
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   627
      then have [simp]: "x [^] (i - j) = \<one>"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   628
        by (simp add: eq assms int_pow_diff)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   629
      then show ?thesis
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   630
        using greater by (rule_tac x="nat (i-j)" in exI) auto
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   631
    qed (use \<open>i \<noteq> j\<close> in auto)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   632
  qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   633
  have 4: "\<exists>i::int. (i \<noteq> 0) \<and> x [^] i = \<one>" if "n \<noteq> 0" "x [^] n = \<one>" for n::nat
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   634
    apply (rule_tac x="int n" in exI)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   635
    by (simp add: int_pow_int that)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   636
  have 5: "finite (carrier (subgroup_generated G {x}))" if "i \<noteq> 0" and 1: "x [^] i = \<one>" for i::int
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   637
  proof -
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   638
    obtain n::nat where n: "n > 0" "x [^] n = \<one>"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   639
      using "1" "3" \<open>i \<noteq> 0\<close> by fastforce
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   640
    have "x [^] a \<in> ([^]) x ` {0..<n}" for a::int
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   641
    proof
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   642
      show "x [^] a = x [^] nat (a mod int n)"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   643
        using n
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   644
        by simp (metis (no_types, lifting) assms dvd_minus_mod dvd_trans int_pow_eq int_pow_eq_id int_pow_int)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   645
      show "nat (a mod int n) \<in> {0..<n}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   646
        using n  apply (simp add:  split: split_nat)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   647
        using Euclidean_Division.pos_mod_bound by presburger
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   648
    qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   649
    then have "carrier (subgroup_generated G {x}) \<subseteq> ([^]) x ` {0..<n}"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   650
      using carrier_subgroup_generated_by_singleton [OF assms] by auto
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
      using finite_surj by blast
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
  show "?fin \<longleftrightarrow> ?nat1" "\<not> ?fin \<longleftrightarrow> ?nateq" "?fin \<longleftrightarrow> ?int1" "\<not> ?fin \<longleftrightarrow> ?inteq"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   655
    using 1 2 3 4 5 by meson+
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   656
qed
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   657
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   658
lemma (in group) finite_cyclic_subgroup_order:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   659
   "x \<in> carrier G \<Longrightarrow> finite(carrier(subgroup_generated G {x})) \<longleftrightarrow> ord x \<noteq> 0"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   660
  by (simp add: finite_cyclic_subgroup ord_eq_0)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   661
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   662
lemma (in group) infinite_cyclic_subgroup_order:
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   663
   "x \<in> carrier G \<Longrightarrow> infinite (carrier(subgroup_generated G {x})) \<longleftrightarrow> ord x = 0"
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   664
  by (simp add: finite_cyclic_subgroup_order)
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   665
733e256ecdf3 new group theory material, mostly ported from HOL Light
paulson <lp15@cam.ac.uk>
parents: 70027
diff changeset
   666
70027
94494b92d8d0 some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
end