| author | wenzelm | 
| Fri, 08 Dec 2023 15:37:46 +0100 | |
| changeset 79207 | f991d3003ec8 | 
| parent 68687 | 2976a4a3b126 | 
| permissions | -rw-r--r-- | 
| 14706 | 1  | 
(* Title: HOL/Algebra/Bij.thy  | 
| 13945 | 2  | 
Author: Florian Kammueller, with new proofs by L C Paulson  | 
3  | 
*)  | 
|
4  | 
||
| 35849 | 5  | 
theory Bij  | 
6  | 
imports Group  | 
|
7  | 
begin  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
16417 
diff
changeset
 | 
8  | 
|
| 61382 | 9  | 
section \<open>Bijections of a Set, Permutation and Automorphism Groups\<close>  | 
| 13945 | 10  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
11  | 
definition  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
12  | 
  Bij :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) set"
 | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67091 
diff
changeset
 | 
13  | 
\<comment> \<open>Only extensional functions, since otherwise we get too many.\<close>  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
14  | 
   where "Bij S = extensional S \<inter> {f. bij_betw f S S}"
 | 
| 13945 | 15  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
16  | 
definition  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
17  | 
  BijGroup :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
18  | 
where "BijGroup S =  | 
| 14963 | 19  | 
\<lparr>carrier = Bij S,  | 
20  | 
mult = \<lambda>g \<in> Bij S. \<lambda>f \<in> Bij S. compose S g f,  | 
|
21  | 
one = \<lambda>x \<in> S. x\<rparr>"  | 
|
| 13945 | 22  | 
|
23  | 
||
24  | 
declare Id_compose [simp] compose_Id [simp]  | 
|
25  | 
||
| 14963 | 26  | 
lemma Bij_imp_extensional: "f \<in> Bij S \<Longrightarrow> f \<in> extensional S"  | 
| 14666 | 27  | 
by (simp add: Bij_def)  | 
| 13945 | 28  | 
|
| 14963 | 29  | 
lemma Bij_imp_funcset: "f \<in> Bij S \<Longrightarrow> f \<in> S \<rightarrow> S"  | 
| 14853 | 30  | 
by (auto simp add: Bij_def bij_betw_imp_funcset)  | 
| 13945 | 31  | 
|
32  | 
||
| 61382 | 33  | 
subsection \<open>Bijections Form a Group\<close>  | 
| 13945 | 34  | 
|
| 33057 | 35  | 
lemma restrict_inv_into_Bij: "f \<in> Bij S \<Longrightarrow> (\<lambda>x \<in> S. (inv_into S f) x) \<in> Bij S"  | 
36  | 
by (simp add: Bij_def bij_betw_inv_into)  | 
|
| 13945 | 37  | 
|
38  | 
lemma id_Bij: "(\<lambda>x\<in>S. x) \<in> Bij S "  | 
|
| 14853 | 39  | 
by (auto simp add: Bij_def bij_betw_def inj_on_def)  | 
| 13945 | 40  | 
|
| 14963 | 41  | 
lemma compose_Bij: "\<lbrakk>x \<in> Bij S; y \<in> Bij S\<rbrakk> \<Longrightarrow> compose S x y \<in> Bij S"  | 
| 14853 | 42  | 
by (auto simp add: Bij_def bij_betw_compose)  | 
| 13945 | 43  | 
|
44  | 
lemma Bij_compose_restrict_eq:  | 
|
| 33057 | 45  | 
"f \<in> Bij S \<Longrightarrow> compose S (restrict (inv_into S f) S) f = (\<lambda>x\<in>S. x)"  | 
46  | 
by (simp add: Bij_def compose_inv_into_id)  | 
|
| 13945 | 47  | 
|
48  | 
theorem group_BijGroup: "group (BijGroup S)"  | 
|
| 68687 | 49  | 
apply (simp add: BijGroup_def)  | 
50  | 
apply (rule groupI)  | 
|
51  | 
apply (auto simp: compose_Bij id_Bij Bij_imp_funcset Bij_imp_extensional compose_assoc [symmetric])  | 
|
52  | 
apply (blast intro: Bij_compose_restrict_eq restrict_inv_into_Bij)  | 
|
53  | 
done  | 
|
| 13945 | 54  | 
|
55  | 
||
| 61382 | 56  | 
subsection\<open>Automorphisms Form a Group\<close>  | 
| 13945 | 57  | 
|
| 33057 | 58  | 
lemma Bij_inv_into_mem: "\<lbrakk> f \<in> Bij S; x \<in> S\<rbrakk> \<Longrightarrow> inv_into S f x \<in> S"  | 
59  | 
by (simp add: Bij_def bij_betw_def inv_into_into)  | 
|
| 13945 | 60  | 
|
| 33057 | 61  | 
lemma Bij_inv_into_lemma:  | 
| 68687 | 62  | 
assumes eq: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> h(g x y) = g (h x) (h y)"  | 
63  | 
and hg: "h \<in> Bij S" "g \<in> S \<rightarrow> S \<rightarrow> S" and "x \<in> S" "y \<in> S"  | 
|
64  | 
shows "inv_into S h (g x y) = g (inv_into S h x) (inv_into S h y)"  | 
|
65  | 
proof -  | 
|
66  | 
have "h ` S = S"  | 
|
67  | 
by (metis (no_types) Bij_def Int_iff assms(2) bij_betw_def mem_Collect_eq)  | 
|
68  | 
with \<open>x \<in> S\<close> \<open>y \<in> S\<close> have "\<exists>x'\<in>S. \<exists>y'\<in>S. x = h x' \<and> y = h y'"  | 
|
69  | 
by auto  | 
|
70  | 
then show ?thesis  | 
|
71  | 
using assms  | 
|
72  | 
by (auto simp add: Bij_def bij_betw_def eq [symmetric] inv_f_f funcset_mem [THEN funcset_mem])  | 
|
73  | 
qed  | 
|
| 13945 | 74  | 
|
| 14963 | 75  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
76  | 
definition  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
77  | 
  auto :: "('a, 'b) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) set"
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
78  | 
where "auto G = hom G G \<inter> Bij (carrier G)"  | 
| 13945 | 79  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
80  | 
definition  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
81  | 
  AutoGroup :: "('a, 'c) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35416 
diff
changeset
 | 
82  | 
where "AutoGroup G = BijGroup (carrier G) \<lparr>carrier := auto G\<rparr>"  | 
| 13945 | 83  | 
|
| 14963 | 84  | 
lemma (in group) id_in_auto: "(\<lambda>x \<in> carrier G. x) \<in> auto G"  | 
| 14666 | 85  | 
by (simp add: auto_def hom_def restrictI group.axioms id_Bij)  | 
| 13945 | 86  | 
|
| 14963 | 87  | 
lemma (in group) mult_funcset: "mult G \<in> carrier G \<rightarrow> carrier G \<rightarrow> carrier G"  | 
| 13945 | 88  | 
by (simp add: Pi_I group.axioms)  | 
89  | 
||
| 33057 | 90  | 
lemma (in group) restrict_inv_into_hom:  | 
| 14963 | 91  | 
"\<lbrakk>h \<in> hom G G; h \<in> Bij (carrier G)\<rbrakk>  | 
| 33057 | 92  | 
\<Longrightarrow> restrict (inv_into (carrier G) h) (carrier G) \<in> hom G G"  | 
93  | 
by (simp add: hom_def Bij_inv_into_mem restrictI mult_funcset  | 
|
94  | 
group.axioms Bij_inv_into_lemma)  | 
|
| 13945 | 95  | 
|
96  | 
lemma inv_BijGroup:  | 
|
| 33057 | 97  | 
"f \<in> Bij S \<Longrightarrow> m_inv (BijGroup S) f = (\<lambda>x \<in> S. (inv_into S f) x)"  | 
| 68687 | 98  | 
apply (rule group.inv_equality [OF group_BijGroup])  | 
| 33057 | 99  | 
apply (simp_all add:BijGroup_def restrict_inv_into_Bij Bij_compose_restrict_eq)  | 
| 13945 | 100  | 
done  | 
101  | 
||
| 14963 | 102  | 
lemma (in group) subgroup_auto:  | 
103  | 
"subgroup (auto G) (BijGroup (carrier G))"  | 
|
104  | 
proof (rule subgroup.intro)  | 
|
105  | 
show "auto G \<subseteq> carrier (BijGroup (carrier G))"  | 
|
106  | 
by (force simp add: auto_def BijGroup_def)  | 
|
107  | 
next  | 
|
108  | 
fix x y  | 
|
109  | 
assume "x \<in> auto G" "y \<in> auto G"  | 
|
110  | 
thus "x \<otimes>\<^bsub>BijGroup (carrier G)\<^esub> y \<in> auto G"  | 
|
111  | 
by (force simp add: BijGroup_def is_group auto_def Bij_imp_funcset  | 
|
112  | 
group.hom_compose compose_Bij)  | 
|
113  | 
next  | 
|
114  | 
show "\<one>\<^bsub>BijGroup (carrier G)\<^esub> \<in> auto G" by (simp add: BijGroup_def id_in_auto)  | 
|
115  | 
next  | 
|
116  | 
fix x  | 
|
117  | 
assume "x \<in> auto G"  | 
|
118  | 
thus "inv\<^bsub>BijGroup (carrier G)\<^esub> x \<in> auto G"  | 
|
119  | 
by (simp del: restrict_apply  | 
|
| 33057 | 120  | 
add: inv_BijGroup auto_def restrict_inv_into_Bij restrict_inv_into_hom)  | 
| 14963 | 121  | 
qed  | 
| 13945 | 122  | 
|
| 14963 | 123  | 
theorem (in group) AutoGroup: "group (AutoGroup G)"  | 
124  | 
by (simp add: AutoGroup_def subgroup.subgroup_is_group subgroup_auto  | 
|
125  | 
group_BijGroup)  | 
|
| 13945 | 126  | 
|
127  | 
end  |