src/HOL/Algebra/Left_Coset.thy
author wenzelm
Fri, 20 Sep 2024 19:51:08 +0200
changeset 80914 d97fdabd9e2b
parent 77089 b4f892d0625d
child 81142 6ad2c917dd2e
permissions -rw-r--r--
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
77089
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
theory Left_Coset
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
imports Coset
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
(*This instance of Coset.thy but for left cosets is due to Jonas Rädle and has been imported
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
  from the AFP entry Orbit_Stabiliser. *)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
begin
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
definition
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 77089
diff changeset
    10
  LCOSETS  :: "[_, 'a set] \<Rightarrow> ('a set)set"   (\<open>lcosets\<index> _\<close> [81] 80)
77089
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
  where "lcosets\<^bsub>G\<^esub> H = (\<Union>a\<in>carrier G. {a <#\<^bsub>G\<^esub> H})"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
definition
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 77089
diff changeset
    14
  LFactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl \<open>LMod\<close> 65)
77089
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
    \<comment> \<open>Actually defined for groups rather than monoids\<close>
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
   where "LFactGroup G H = \<lparr>carrier = lcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
lemma (in group) lcos_self: "[| x \<in> carrier G; subgroup H G |] ==> x \<in> x <# H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
  by (simp add: group_l_invI subgroup.lcos_module_rev subgroup.one_closed)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
text \<open>Elements of a left coset are in the carrier\<close>
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
lemma (in subgroup) elemlcos_carrier:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
  assumes "group G" "a \<in> carrier G" "a' \<in> a <# H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
  shows "a' \<in> carrier G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  by (meson assms group.l_coset_carrier subgroup_axioms)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
text \<open>Step one for lemma \<open>rcos_module\<close>\<close>
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
lemma (in subgroup) lcos_module_imp:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
  assumes "group G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  assumes xcarr: "x \<in> carrier G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
      and x'cos: "x' \<in> x <# H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
  shows "(inv x \<otimes> x') \<in> H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
proof -
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
  interpret group G by fact
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
  obtain h
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
    where hH: "h \<in> H" and x': "x' = x \<otimes> h" and hcarr: "h \<in> carrier G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
    using assms by (auto simp: l_coset_def)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
  have "(inv x) \<otimes> x' = (inv x) \<otimes> (x \<otimes> h)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
    by (simp add: x')
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
  have "\<dots> = (x \<otimes> inv x) \<otimes> h"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
    by (metis hcarr inv_closed inv_inv l_inv m_assoc xcarr)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  also have "\<dots> = h"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
    by (simp add: hcarr xcarr)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
  finally have "(inv x) \<otimes> x' = h"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
    using x' by metis
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
  then show "(inv x) \<otimes> x' \<in> H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
    using hH by force
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
text \<open>Left cosets are subsets of the carrier.\<close> 
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
lemma (in subgroup) lcosets_carrier:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
  assumes "group G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
  assumes XH: "X \<in> lcosets H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
  shows "X \<subseteq> carrier G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
proof -
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
  interpret group G by fact
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
  show "X \<subseteq> carrier G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
    using XH l_coset_subset_G subset by (auto simp: LCOSETS_def)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
lemma (in group) lcosets_part_G:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
  assumes "subgroup H G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  shows "\<Union>(lcosets H) = carrier G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
proof -
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
  interpret subgroup H G by fact
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  show ?thesis
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  proof
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
    show "\<Union> (lcosets H) \<subseteq> carrier G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
      by (force simp add: LCOSETS_def l_coset_def)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
    show "carrier G \<subseteq> \<Union> (lcosets H)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
      by (auto simp add: LCOSETS_def intro: lcos_self assms)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
  qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
lemma (in group) lcosets_subset_PowG:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
     "subgroup H G  \<Longrightarrow> lcosets H \<subseteq> Pow(carrier G)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
  using lcosets_part_G subset_Pow_Union by blast
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
lemma (in group) lcos_disjoint:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  assumes "subgroup H G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
  assumes p: "a \<in> lcosets H" "b \<in> lcosets H" "a\<noteq>b"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
  shows "a \<inter> b = {}"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
proof -
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
  interpret subgroup H G by fact
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  show ?thesis
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
    using p l_repr_independence subgroup_axioms unfolding LCOSETS_def disjoint_iff
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
    by blast
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
text\<open>The next two lemmas support the proof of \<open>card_cosets_equal\<close>.\<close>
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
lemma (in group) inj_on_f':
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
    "\<lbrakk>H \<subseteq> carrier G;  a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. y \<otimes> inv a) (a <# H)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
  by (simp add: inj_on_g l_coset_subset_G)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
lemma (in group) inj_on_f'':
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
    "\<lbrakk>H \<subseteq> carrier G;  a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. inv a \<otimes> y) (a <# H)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
  by (meson inj_on_cmult inv_closed l_coset_subset_G subset_inj_on)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
lemma (in group) inj_on_g':
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
    "\<lbrakk>H \<subseteq> carrier G;  a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. a \<otimes> y) H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  using inj_on_cmult inj_on_subset by blast
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
lemma (in group) l_card_cosets_equal:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  assumes "c \<in> lcosets H" and H: "H \<subseteq> carrier G" and fin: "finite(carrier G)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
  shows "card H = card c"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
proof -
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
  obtain x where x: "x \<in> carrier G" and c: "c = x <# H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
    using assms by (auto simp add: LCOSETS_def)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
  have "inj_on ((\<otimes>) x) H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
    by (simp add: H group.inj_on_g' x)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
  moreover
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
  have "(\<otimes>) x ` H \<subseteq> x <# H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
    by (force simp add: m_assoc subsetD l_coset_def)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  moreover
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
  have "inj_on ((\<otimes>) (inv x)) (x <# H)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    by (simp add: H group.inj_on_f'' x)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
  moreover
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
  have "\<And>h. h \<in> H \<Longrightarrow> inv x \<otimes> (x \<otimes> h) \<in> H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
    by (metis H in_mono inv_solve_left m_closed x)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  then have "(\<otimes>) (inv x) ` (x <# H) \<subseteq> H"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
    by (auto simp: l_coset_def)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
  ultimately show ?thesis
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
    by (metis H fin c card_bij_eq finite_imageD finite_subset)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
theorem (in group) l_lagrange:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
  assumes "finite(carrier G)" "subgroup H G"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  shows "card(lcosets H) * card(H) = order(G)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
proof -
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
  have "card H * card (lcosets H) = card (\<Union> (lcosets H))"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
    using card_partition
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
    by (metis (no_types, lifting) assms finite_UnionD l_card_cosets_equal lcos_disjoint lcosets_part_G subgroup.subset)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
  then show ?thesis
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
    by (simp add: assms(2) lcosets_part_G mult.commute order_def)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
end