| author | wenzelm | 
| Fri, 24 Nov 2023 20:58:12 +0100 | |
| changeset 79053 | badb3da19ac6 | 
| parent 73648 | 1bd3463e30b8 | 
| child 81438 | 95c9af7483b1 | 
| permissions | -rw-r--r-- | 
| 68582 | 1 | (* Title: HOL/Algebra/Sym_Groups.thy | 
| 2 | Author: Paulo Emílio de Vilhena | |
| 3 | *) | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | theory Sym_Groups | 
| 73477 | 6 | imports | 
| 7 | "HOL-Combinatorics.Cycles" | |
| 8 | Solvable_Groups | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | begin | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 11 | section \<open>Symmetric Groups\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 12 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 13 | subsection \<open>Definitions\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 14 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | abbreviation inv' :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | where "inv' f \<equiv> Hilbert_Choice.inv f" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | definition sym_group :: "nat \<Rightarrow> (nat \<Rightarrow> nat) monoid" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 |   where "sym_group n = \<lparr> carrier = { p. p permutes {1..n} }, mult = (\<circ>), one = id \<rparr>"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 21 | definition alt_group :: "nat \<Rightarrow> (nat \<Rightarrow> nat) monoid" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 22 |   where "alt_group n = (sym_group n) \<lparr> carrier := { p. p permutes {1..n} \<and> evenperm p } \<rparr>"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 23 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | definition sign_img :: "int monoid" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68975diff
changeset | 25 |   where "sign_img = \<lparr> carrier = { -1, 1 }, mult = (*), one = 1 \<rparr>"
 | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 28 | subsection \<open>Basic Properties\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 29 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 30 | lemma sym_group_carrier: "p \<in> carrier (sym_group n) \<longleftrightarrow> p permutes {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 31 | unfolding sym_group_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 32 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 33 | lemma sym_group_mult: "mult (sym_group n) = (\<circ>)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 34 | unfolding sym_group_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 35 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 36 | lemma sym_group_one: "one (sym_group n) = id" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 37 | unfolding sym_group_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 38 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 39 | lemma sym_group_carrier': "p \<in> carrier (sym_group n) \<Longrightarrow> permutation p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 40 | unfolding sym_group_carrier permutation_permutes by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 41 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 42 | lemma alt_group_carrier: "p \<in> carrier (alt_group n) \<longleftrightarrow> p permutes {1..n} \<and> evenperm p"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 43 | unfolding alt_group_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 44 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 45 | lemma alt_group_mult: "mult (alt_group n) = (\<circ>)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 46 | unfolding alt_group_def using sym_group_mult by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 47 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 48 | lemma alt_group_one: "one (alt_group n) = id" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 49 | unfolding alt_group_def using sym_group_one by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 50 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 51 | lemma alt_group_carrier': "p \<in> carrier (alt_group n) \<Longrightarrow> permutation p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 52 | unfolding alt_group_carrier permutation_permutes by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 53 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | lemma sym_group_is_group: "group (sym_group n)" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 55 | using permutes_inv permutes_inv_o(2) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 56 | by (auto intro!: groupI | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 57 | simp add: sym_group_def permutes_compose | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 58 | permutes_id comp_assoc, blast) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 59 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 60 | lemma sign_img_is_group: "group sign_img" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 61 | unfolding sign_img_def by (unfold_locales, auto simp add: Units_def) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | lemma sym_group_inv_closed: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 64 | assumes "p \<in> carrier (sym_group n)" shows "inv' p \<in> carrier (sym_group n)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | using assms permutes_inv sym_group_def by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 67 | lemma alt_group_inv_closed: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 68 | assumes "p \<in> carrier (alt_group n)" shows "inv' p \<in> carrier (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 69 | using evenperm_inv[OF alt_group_carrier'] permutes_inv assms alt_group_carrier by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 70 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 71 | lemma sym_group_inv_equality [simp]: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 72 | assumes "p \<in> carrier (sym_group n)" shows "inv\<^bsub>(sym_group n)\<^esub> p = inv' p" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | proof - | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | have "inv' p \<circ> p = id" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | using assms permutes_inv_o(2) sym_group_def by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | hence "(inv' p) \<otimes>\<^bsub>(sym_group n)\<^esub> p = one (sym_group n)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | by (simp add: sym_group_def) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | thus ?thesis using group.inv_equality[OF sym_group_is_group] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | by (simp add: assms sym_group_inv_closed) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 82 | lemma sign_is_hom: "sign \<in> hom (sym_group n) sign_img" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 83 | unfolding hom_def sign_img_def sym_group_mult using sym_group_carrier'[of _ n] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 84 | by (auto simp add: sign_compose, meson sign_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 85 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 86 | lemma sign_group_hom: "group_hom (sym_group n) sign_img sign" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 87 | using group_hom.intro[OF sym_group_is_group sign_img_is_group] sign_is_hom | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 88 | by (simp add: group_hom_axioms_def) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 90 | lemma sign_is_surj: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 91 | assumes "n \<ge> 2" shows "sign ` (carrier (sym_group n)) = carrier sign_img" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 92 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 93 | have "swapidseq (Suc 0) (Fun.swap (1 :: nat) 2 id)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 94 | using comp_Suc[OF id, of "1 :: nat" "2"] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 95 | hence "sign (Fun.swap (1 :: nat) 2 id) = (-1 :: int)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 96 | by (simp add: sign_swap_id) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 97 | moreover have "Fun.swap (1 :: nat) 2 id \<in> carrier (sym_group n)" and "id \<in> carrier (sym_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 98 |     using assms permutes_swap_id[of "1 :: nat" "{1..n}" 2] permutes_id
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 99 | unfolding sym_group_carrier by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 100 | ultimately have "carrier sign_img \<subseteq> sign ` (carrier (sym_group n))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 101 | using sign_id mk_disjoint_insert unfolding sign_img_def by fastforce | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 102 | moreover have "sign ` (carrier (sym_group n)) \<subseteq> carrier sign_img" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 103 | using sign_is_hom unfolding hom_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 104 | ultimately show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 105 | by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 106 | qed | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 108 | lemma alt_group_is_sign_kernel: | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | "carrier (alt_group n) = kernel (sym_group n) sign_img sign" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | unfolding alt_group_def sym_group_def sign_img_def kernel_def sign_def by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 112 | lemma alt_group_is_subgroup: "subgroup (carrier (alt_group n)) (sym_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 113 | using group_hom.subgroup_kernel[OF sign_group_hom] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 114 | unfolding alt_group_is_sign_kernel by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 116 | lemma alt_group_is_group: "group (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 117 | using group.subgroup_imp_group[OF sym_group_is_group alt_group_is_subgroup] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 118 | by (simp add: alt_group_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 119 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 120 | lemma sign_iso: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 121 | assumes "n \<ge> 2" shows "(sym_group n) Mod (carrier (alt_group n)) \<cong> sign_img" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 122 | using group_hom.FactGroup_iso[OF sign_group_hom sign_is_surj[OF assms]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 123 | unfolding alt_group_is_sign_kernel . | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | lemma alt_group_inv_equality: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 126 | assumes "p \<in> carrier (alt_group n)" shows "inv\<^bsub>(alt_group n)\<^esub> p = inv' p" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | proof - | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | have "inv' p \<circ> p = id" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 | using assms permutes_inv_o(2) alt_group_def by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | hence "(inv' p) \<otimes>\<^bsub>(alt_group n)\<^esub> p = one (alt_group n)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | by (simp add: alt_group_def sym_group_def) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | thus ?thesis using group.inv_equality[OF alt_group_is_group] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | by (simp add: assms alt_group_inv_closed) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 136 | lemma sym_group_card_carrier: "card (carrier (sym_group n)) = fact n" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 137 |   using card_permutations[of "{1..n}" n] unfolding sym_group_def by simp
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 138 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 139 | lemma alt_group_card_carrier: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 140 | assumes "n \<ge> 2" shows "2 * card (carrier (alt_group n)) = fact n" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 141 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 142 | have "card (rcosets\<^bsub>sym_group n\<^esub> (carrier (alt_group n))) = 2" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 143 | using iso_same_card[OF sign_iso[OF assms]] unfolding FactGroup_def sign_img_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 144 | thus ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 145 | using group.lagrange[OF sym_group_is_group alt_group_is_subgroup, of n] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 146 | unfolding order_def sym_group_card_carrier by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 147 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 148 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | subsection \<open>Transposition Sequences\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | text \<open>In order to prove that the Alternating Group can be generated by 3-cycles, we need | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | a stronger decomposition of permutations as transposition sequences than the one | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 154 | proposed at Permutations.thy. \<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 156 | inductive swapidseq_ext :: "'a set \<Rightarrow> nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 157 | where | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 158 |     empty:  "swapidseq_ext {} 0 id"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 159 | | single: "\<lbrakk> swapidseq_ext S n p; a \<notin> S \<rbrakk> \<Longrightarrow> swapidseq_ext (insert a S) n p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 160 | | comp: "\<lbrakk> swapidseq_ext S n p; a \<noteq> b \<rbrakk> \<Longrightarrow> | 
| 73648 | 161 | swapidseq_ext (insert a (insert b S)) (Suc n) ((transpose a b) \<circ> p)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | lemma swapidseq_ext_finite: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 165 | assumes "swapidseq_ext S n p" shows "finite S" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 166 | using assms by (induction) (auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 167 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 168 | lemma swapidseq_ext_zero: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 169 | assumes "finite S" shows "swapidseq_ext S 0 id" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 170 | using assms empty by (induct set: "finite", fastforce, simp add: single) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 171 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 172 | lemma swapidseq_ext_imp_swapidseq: | 
| 73477 | 173 | \<open>swapidseq n p\<close> if \<open>swapidseq_ext S n p\<close> | 
| 174 | using that proof induction | |
| 175 | case empty | |
| 176 | then show ?case | |
| 177 | by (simp add: fun_eq_iff) | |
| 178 | next | |
| 179 | case (single S n p a) | |
| 180 | then show ?case by simp | |
| 181 | next | |
| 182 | case (comp S n p a b) | |
| 73648 | 183 | then have \<open>swapidseq (Suc n) (transpose a b \<circ> p)\<close> | 
| 73477 | 184 | by (simp add: comp_Suc) | 
| 185 | then show ?case by (simp add: comp_def) | |
| 186 | qed | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | lemma swapidseq_ext_zero_imp_id: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 189 | assumes "swapidseq_ext S 0 p" shows "p = id" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 191 | have "\<lbrakk> swapidseq_ext S n p; n = 0 \<rbrakk> \<Longrightarrow> p = id" for n | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 192 | by (induction rule: swapidseq_ext.induct, auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 193 | thus ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 194 | using assms by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | lemma swapidseq_ext_finite_expansion: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 198 | assumes "finite B" and "swapidseq_ext A n p" shows "swapidseq_ext (A \<union> B) n p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 199 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 200 | proof (induct set: "finite", simp) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 201 | case (insert b B) show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 202 | using insert single[OF insert(3), of b] by (metis Un_insert_right insert_absorb) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | lemma swapidseq_ext_backwards: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | assumes "swapidseq_ext A (Suc n) p" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | shows "\<exists>a b A' p'. a \<noteq> b \<and> A = (insert a (insert b A')) \<and> | 
| 73648 | 208 | swapidseq_ext A' n p' \<and> p = (transpose a b) \<circ> p'" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 210 |   { fix A n k and p :: "'a \<Rightarrow> 'a"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 211 | assume "swapidseq_ext A n p" "n = Suc k" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | hence "\<exists>a b A' p'. a \<noteq> b \<and> A = (insert a (insert b A')) \<and> | 
| 73648 | 213 | swapidseq_ext A' k p' \<and> p = (transpose a b) \<circ> p'" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 214 | proof (induction, simp) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | case single thus ?case | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | by (metis Un_insert_right insert_iff insert_is_Un swapidseq_ext.single) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 218 | case comp thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 219 | by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | qed } | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 221 | thus ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 222 | using assms by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 225 | lemma swapidseq_ext_backwards': | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 226 | assumes "swapidseq_ext A (Suc n) p" | 
| 73648 | 227 | shows "\<exists>a b A' p'. a \<in> A \<and> b \<in> A \<and> a \<noteq> b \<and> swapidseq_ext A n p' \<and> p = (transpose a b) \<circ> p'" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 228 | using swapidseq_ext_backwards[OF assms] swapidseq_ext_finite_expansion | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 229 | by (metis Un_insert_left assms insertI1 sup.idem sup_commute swapidseq_ext_finite) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 230 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | lemma swapidseq_ext_endswap: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | assumes "swapidseq_ext S n p" "a \<noteq> b" | 
| 73648 | 233 | shows "swapidseq_ext (insert a (insert b S)) (Suc n) (p \<circ> (transpose a b))" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 234 | using assms | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | proof (induction n arbitrary: S p a b) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | case 0 hence "p = id" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | using swapidseq_ext_zero_imp_id by blast | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 238 | thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 239 | using 0 by (metis comp_id id_comp swapidseq_ext.comp) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | case (Suc n) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | then obtain c d S' and p' :: "'a \<Rightarrow> 'a" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 243 | where cd: "c \<noteq> d" and S: "S = (insert c (insert d S'))" "swapidseq_ext S' n p'" | 
| 73648 | 244 | and p: "p = transpose c d \<circ> p'" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | using swapidseq_ext_backwards[OF Suc(2)] by blast | 
| 73648 | 246 | hence "swapidseq_ext (insert a (insert b S')) (Suc n) (p' \<circ> (transpose a b))" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | by (simp add: Suc.IH Suc.prems(2)) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | hence "swapidseq_ext (insert c (insert d (insert a (insert b S')))) (Suc (Suc n)) | 
| 73648 | 249 | (transpose c d \<circ> p' \<circ> (transpose a b))" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | by (metis cd fun.map_comp swapidseq_ext.comp) | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 251 | thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 252 | by (metis S(1) p insert_commute) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | lemma swapidseq_ext_extension: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 256 |   assumes "swapidseq_ext A n p" and "swapidseq_ext B m q" and "A \<inter> B = {}"
 | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | shows "swapidseq_ext (A \<union> B) (n + m) (p \<circ> q)" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 258 | using assms(1,3) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 259 | proof (induction, simp add: assms(2)) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 260 | case single show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 261 | using swapidseq_ext.single[OF single(3)] single(2,4) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 262 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 263 | case comp show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 264 | using swapidseq_ext.comp[OF comp(3,2)] comp(4) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 265 | by (metis Un_insert_left add_Suc insert_disjoint(1) o_assoc) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | lemma swapidseq_ext_of_cycles: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 269 | assumes "cycle cs" shows "swapidseq_ext (set cs) (length cs - 1) (cycle_of_list cs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 270 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 271 | proof (induct cs rule: cycle_of_list.induct) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | case (1 i j cs) show ?case | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 273 | using comp[OF 1(1), of i j] 1(2) by (simp add: o_def) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 275 | case "2_1" show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 276 | by (simp, metis eq_id_iff empty) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 278 |   case ("2_2" v) show ?case
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 279 | using single[OF empty, of v] by (simp, metis eq_id_iff) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | lemma cycle_decomp_imp_swapidseq_ext: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 283 | assumes "cycle_decomp S p" shows "\<exists>n. swapidseq_ext S n p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 284 | using assms | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | proof (induction) | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 286 | case empty show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 287 | using swapidseq_ext.empty by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | case (comp I p cs) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | then obtain m where m: "swapidseq_ext I m p" by blast | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | hence "swapidseq_ext (set cs) (length cs - 1) (cycle_of_list cs)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | using comp.hyps(2) swapidseq_ext_of_cycles by blast | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | thus ?case using swapidseq_ext_extension m | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | using comp.hyps(3) by blast | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 297 | lemma swapidseq_ext_of_permutation: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 298 | assumes "p permutes S" and "finite S" shows "\<exists>n. swapidseq_ext S n p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 299 | using cycle_decomp_imp_swapidseq_ext[OF cycle_decomposition[OF assms]] . | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | lemma split_swapidseq_ext: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 302 | assumes "k \<le> n" and "swapidseq_ext S n p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 303 | obtains q r U V where "swapidseq_ext U (n - k) q" and "swapidseq_ext V k r" and "p = q \<circ> r" and "U \<union> V = S" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 304 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 305 | from assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 306 | have "\<exists>q r U V. swapidseq_ext U (n - k) q \<and> swapidseq_ext V k r \<and> p = q \<circ> r \<and> U \<union> V = S" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 307 | (is "\<exists>q r U V. ?split k q r U V") | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 308 | proof (induct k rule: inc_induct) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 309 | case base thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 310 | by (metis diff_self_eq_0 id_o sup_bot.left_neutral empty) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 311 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 312 | case (step m) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 313 | then obtain q r U V | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 314 | where q: "swapidseq_ext U (n - Suc m) q" and r: "swapidseq_ext V (Suc m) r" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 315 | and p: "p = q \<circ> r" and S: "U \<union> V = S" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 316 | by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 317 | obtain a b r' V' | 
| 73648 | 318 | where "a \<noteq> b" and r': "V = (insert a (insert b V'))" "swapidseq_ext V' m r'" "r = (transpose a b) \<circ> r'" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 319 | using swapidseq_ext_backwards[OF r] by blast | 
| 73648 | 320 | have "swapidseq_ext (insert a (insert b U)) (n - m) (q \<circ> (transpose a b))" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 321 | using swapidseq_ext_endswap[OF q \<open>a \<noteq> b\<close>] step(2) by (metis Suc_diff_Suc) | 
| 73648 | 322 | hence "?split m (q \<circ> (transpose a b)) r' (insert a (insert b U)) V'" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 323 | using r' S unfolding p by fastforce | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 324 | thus ?case by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 325 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 326 | thus ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 327 | using that by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 331 | subsection \<open>Unsolvability of Symmetric Groups\<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | |
| 69597 | 333 | text \<open>We show that symmetric groups (\<^term>\<open>sym_group n\<close>) are unsolvable for \<^term>\<open>n \<ge> 5\<close>.\<close> | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 334 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 335 | abbreviation three_cycles :: "nat \<Rightarrow> (nat \<Rightarrow> nat) set" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | where "three_cycles n \<equiv> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 |            { cycle_of_list cs | cs. cycle cs \<and> length cs = 3 \<and> set cs \<subseteq> {1..n} }"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | lemma stupid_lemma: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 341 | assumes "length cs = 3" shows "cs = [ (cs ! 0), (cs ! 1), (cs ! 2) ]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 342 | using assms by (auto intro!: nth_equalityI) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 343 | (metis Suc_lessI less_2_cases not_less_eq nth_Cons_0 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 344 | nth_Cons_Suc numeral_2_eq_2 numeral_3_eq_3) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 345 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 346 | lemma three_cycles_incl: "three_cycles n \<subseteq> carrier (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 347 | proof | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 348 | fix p assume "p \<in> three_cycles n" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 349 |   then obtain cs where cs: "p = cycle_of_list cs" "cycle cs" "length cs = 3" "set cs \<subseteq> {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 350 | by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 351 | obtain a b c where cs_def: "cs = [ a, b, c ]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 352 | using stupid_lemma[OF cs(3)] by auto | 
| 73648 | 353 | have "swapidseq (Suc (Suc 0)) ((transpose a b) \<circ> (Fun.swap b c id))" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 354 | using comp_Suc[OF comp_Suc[OF id], of b c a b] cs(2) unfolding cs_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 355 | hence "evenperm p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 356 | using cs(1) unfolding cs_def by (simp add: evenperm_unique) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 357 | thus "p \<in> carrier (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 358 | using permutes_subset[OF cycle_permutes cs(4)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 359 | unfolding alt_group_carrier cs(1) by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 362 | lemma alt_group_carrier_as_three_cycles: | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | "carrier (alt_group n) = generate (alt_group n) (three_cycles n)" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 364 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 365 | interpret A: group "alt_group n" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 366 | using alt_group_is_group by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 368 | show ?thesis | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | proof | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 370 | show "generate (alt_group n) (three_cycles n) \<subseteq> carrier (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 371 | using A.generate_subgroup_incl[OF three_cycles_incl A.subgroup_self] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 372 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 373 | show "carrier (alt_group n) \<subseteq> generate (alt_group n) (three_cycles n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 374 | proof | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 |       { fix q :: "nat \<Rightarrow> nat" and a b c
 | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 376 |         assume "a \<noteq> b" "b \<noteq> c" "{ a, b, c } \<subseteq> {1..n}" 
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 377 | have "cycle_of_list [a, b, c] \<in> generate (alt_group n) (three_cycles n)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | proof (cases) | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 379 | assume "c = a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 380 | hence "cycle_of_list [ a, b, c ] = id" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 381 | by (simp add: swap_commute) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 382 | thus "cycle_of_list [ a, b, c ] \<in> generate (alt_group n) (three_cycles n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 383 | using one[of "alt_group n"] unfolding alt_group_one by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 385 | assume "c \<noteq> a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 386 | have "distinct [a, b, c]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 387 | using \<open>a \<noteq> b\<close> and \<open>b \<noteq> c\<close> and \<open>c \<noteq> a\<close> by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 388 |           with \<open>{ a, b, c } \<subseteq> {1..n}\<close>
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 389 | show "cycle_of_list [ a, b, c ] \<in> generate (alt_group n) (three_cycles n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 390 | by (intro incl, fastforce) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 391 | qed } note aux_lemma1 = this | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 392 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 393 |       { fix S :: "nat set" and q :: "nat \<Rightarrow> nat"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 394 |         assume seq: "swapidseq_ext S (Suc (Suc 0)) q" and S: "S \<subseteq> {1..n}"
 | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | have "q \<in> generate (alt_group n) (three_cycles n)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 397 | obtain a b q' where ab: "a \<noteq> b" "a \<in> S" "b \<in> S" | 
| 73648 | 398 | and q': "swapidseq_ext S (Suc 0) q'" "q = (transpose a b) \<circ> q'" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 399 | using swapidseq_ext_backwards'[OF seq] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 400 | obtain c d where cd: "c \<noteq> d" "c \<in> S" "d \<in> S" | 
| 73648 | 401 | and q: "q = (transpose a b) \<circ> (Fun.swap c d id)" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 402 | using swapidseq_ext_backwards'[OF q'(1)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 403 | swapidseq_ext_zero_imp_id | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 404 | unfolding q'(2) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 405 | by fastforce | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 406 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 407 | consider (eq) "b = c" | (ineq) "b \<noteq> c" by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | thus ?thesis | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 409 | proof cases | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 410 | case eq then have "q = cycle_of_list [ a, b, d ]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 411 | unfolding q by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 412 |             moreover have "{ a, b, d } \<subseteq> {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 413 | using ab cd S by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 414 | ultimately show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 415 | using aux_lemma1[OF ab(1)] cd(1) eq by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 417 | case ineq | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 418 | hence "q = cycle_of_list [ a, b, c ] \<circ> cycle_of_list [ b, c, d ]" | 
| 73648 | 419 | unfolding q by (simp add: swap_nilpotent o_assoc) | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 420 |             moreover have "{ a, b, c } \<subseteq> {1..n}" and "{ b, c, d } \<subseteq> {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 421 | using ab cd S by blast+ | 
| 68604 | 422 | ultimately show ?thesis | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 423 | using eng[OF aux_lemma1[OF ab(1) ineq] aux_lemma1[OF ineq cd(1)]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 424 | unfolding alt_group_mult by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | qed | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 426 | qed } note aux_lemma2 = this | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 428 |       fix p assume "p \<in> carrier (alt_group n)" then have p: "p permutes {1..n}" "evenperm p"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 429 | unfolding alt_group_carrier by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 430 |       obtain m where m: "swapidseq_ext {1..n} m p"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 431 | using swapidseq_ext_of_permutation[OF p(1)] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 432 | have "even m" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 433 | using swapidseq_ext_imp_swapidseq[OF m] p(2) evenperm_unique by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 434 | then obtain k where k: "m = 2 * k" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 435 | by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 436 | show "p \<in> generate (alt_group n) (three_cycles n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 437 | using m unfolding k | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | proof (induct k arbitrary: p) | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 439 | case 0 then have "p = id" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 440 | using swapidseq_ext_zero_imp_id by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 441 | show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 442 | using generate.one[of "alt_group n" "three_cycles n"] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 443 | unfolding alt_group_one \<open>p = id\<close> . | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 445 | case (Suc m) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 446 | have arith: "2 * (Suc m) - (Suc (Suc 0)) = 2 * m" and "Suc (Suc 0) \<le> 2 * Suc m" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 447 | by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 448 | then obtain q r U V | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 449 | where q: "swapidseq_ext U (2 * m) q" and r: "swapidseq_ext V (Suc (Suc 0)) r" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 450 |             and p: "p = q \<circ> r" and UV: "U \<union> V = {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 451 | using split_swapidseq_ext[OF _ Suc(2), of "Suc (Suc 0)"] unfolding arith by metis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 452 |         have "swapidseq_ext {1..n} (2 * m) q"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 453 | using UV q swapidseq_ext_finite_expansion[OF swapidseq_ext_finite[OF r] q] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 454 | thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 455 | using eng[OF Suc(1) aux_lemma2[OF r], of q] UV unfolding alt_group_mult p by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 456 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 459 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 460 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 461 | theorem derived_alt_group_const: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 462 | assumes "n \<ge> 5" shows "derived (alt_group n) (carrier (alt_group n)) = carrier (alt_group n)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 463 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 | show "derived (alt_group n) (carrier (alt_group n)) \<subseteq> carrier (alt_group n)" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 465 | using group.derived_in_carrier[OF alt_group_is_group] by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 467 |   { fix p assume "p \<in> three_cycles n" have "p \<in> derived (alt_group n) (carrier (alt_group n))"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 468 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 469 |       obtain cs where cs: "p = cycle_of_list cs" "cycle cs" "length cs = 3" "set cs \<subseteq> {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 470 | using \<open>p \<in> three_cycles n\<close> by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 471 | then obtain a b c where cs_def: "cs = [ a, b, c ]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 472 | using stupid_lemma[OF cs(3)] by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 473 | have "card (set cs) = 3" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 474 | using cs(2-3) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 475 | by (simp add: distinct_card) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 477 |       have "set cs \<noteq> {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 478 | using assms cs(3) unfolding sym[OF distinct_card[OF cs(2)]] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 479 |       then obtain d where d: "d \<notin> set cs" "d \<in> {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 480 | using cs(4) by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 481 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 482 |       hence "cycle (d # cs)" and "length (d # cs) = 4" and "card {1..n} = n"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 483 | using cs(2-3) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 484 |       hence "set (d # cs) \<noteq> {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 485 | using assms unfolding sym[OF distinct_card[OF \<open>cycle (d # cs)\<close>]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 486 | by (metis Suc_n_not_le_n eval_nat_numeral(3)) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 487 |       then obtain e where e: "e \<notin> set (d # cs)" "e \<in> {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 488 | using d cs(4) by (metis insert_subset list.simps(15) subsetI subset_antisym) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 489 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 490 | define q where "q = (Fun.swap d e id) \<circ> (Fun.swap b c id)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 491 | hence "bij q" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 492 | by (simp add: bij_comp) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 493 | moreover have "q b = c" and "q c = b" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 494 | using d(1) e(1) unfolding q_def cs_def by simp+ | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 495 | moreover have "q a = a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 496 | using d(1) e(1) cs(2) unfolding q_def cs_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 497 | ultimately have "q \<circ> p \<circ> (inv' q) = cycle_of_list [ a, c, b ]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 498 | using conjugation_of_cycle[OF cs(2), of q] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 499 | unfolding sym[OF cs(1)] unfolding cs_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 500 | also have " ... = p \<circ> p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 501 | using cs(2) unfolding cs(1) cs_def | 
| 73648 | 502 | by (simp add: comp_swap swap_commute transpose_comp_triple) | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 503 | finally have "q \<circ> p \<circ> (inv' q) = p \<circ> p" . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 504 | moreover have "bij p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 505 | unfolding cs(1) cs_def by (simp add: bij_comp) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 506 | ultimately have p: "q \<circ> p \<circ> (inv' q) \<circ> (inv' p) = p" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 507 | by (simp add: bijection.intro bijection.inv_comp_right comp_assoc) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 508 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 509 | have "swapidseq (Suc (Suc 0)) q" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 510 | using comp_Suc[OF comp_Suc[OF id], of b c d e] e(1) cs(2) unfolding q_def cs_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 511 | hence "evenperm q" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 512 | using even_Suc_Suc_iff evenperm_unique by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 513 |       moreover have "q permutes { d, e, b, c }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 514 | unfolding q_def by (simp add: permutes_compose permutes_swap_id) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 515 |       hence "q permutes {1..n}"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 516 | using cs(4) d(2) e(2) permutes_subset unfolding cs_def by fastforce | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 517 | ultimately have "q \<in> carrier (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 518 | unfolding alt_group_carrier by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 519 | moreover have "p \<in> carrier (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 520 | using \<open>p \<in> three_cycles n\<close> three_cycles_incl by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 521 | ultimately have "p \<in> derived_set (alt_group n) (carrier (alt_group n))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 522 | using p alt_group_inv_equality unfolding alt_group_mult | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 523 | by (metis (no_types, lifting) UN_iff singletonI) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 524 | thus "p \<in> derived (alt_group n) (carrier (alt_group n))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 525 | unfolding derived_def by (rule incl) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 526 | qed } note aux_lemma = this | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 527 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 528 | interpret A: group "alt_group n" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 529 | using alt_group_is_group . | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 530 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 531 | have "generate (alt_group n) (three_cycles n) \<subseteq> derived (alt_group n) (carrier (alt_group n))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 532 | using A.generate_subgroup_incl[OF _ A.derived_is_subgroup] aux_lemma by (meson subsetI) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 533 | thus "carrier (alt_group n) \<subseteq> derived (alt_group n) (carrier (alt_group n))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 534 | using alt_group_carrier_as_three_cycles by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 535 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 536 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 537 | corollary alt_group_is_unsolvable: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 538 | assumes "n \<ge> 5" shows "\<not> solvable (alt_group n)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 539 | proof (rule ccontr) | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 540 | assume "\<not> \<not> solvable (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 541 |   then obtain m where "(derived (alt_group n) ^^ m) (carrier (alt_group n)) = { id }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 542 | using group.solvable_iff_trivial_derived_seq[OF alt_group_is_group] unfolding alt_group_one by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 543 | moreover have "(derived (alt_group n) ^^ m) (carrier (alt_group n)) = carrier (alt_group n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 544 | using derived_alt_group_const[OF assms] by (induct m) (auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 545 | ultimately have card_eq_1: "card (carrier (alt_group n)) = 1" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 546 | by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 547 | have ge_2: "n \<ge> 2" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 548 | using assms by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 549 | moreover have "2 = fact n" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 550 | using alt_group_card_carrier[OF ge_2] unfolding card_eq_1 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 551 | by (metis fact_2 mult.right_neutral of_nat_fact) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 552 | ultimately have "n = 2" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 553 | by (metis antisym_conv fact_ge_self) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 554 | thus False | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 555 | using assms by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 556 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 557 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 558 | corollary sym_group_is_unsolvable: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 559 | assumes "n \<ge> 5" shows "\<not> solvable (sym_group n)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 561 | interpret Id: group_hom "alt_group n" "sym_group n" id | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 562 | using group.canonical_inj_is_hom[OF sym_group_is_group alt_group_is_subgroup] alt_group_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 563 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 564 | using Id.inj_hom_imp_solvable alt_group_is_unsolvable[OF assms] by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 565 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 567 | end |