| 
14706
 | 
     1  | 
(*  Title:      HOL/Algebra/Bij.thy
  | 
| 
13945
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Author:     Florian Kammueller, with new proofs by L C Paulson
  | 
| 
 | 
     4  | 
*)
  | 
| 
 | 
     5  | 
  | 
| 
14666
 | 
     6  | 
header {* Bijections of a Set, Permutation Groups, Automorphism Groups *}
 | 
| 
13945
 | 
     7  | 
  | 
| 
16417
 | 
     8  | 
theory Bij imports Group begin
  | 
| 
13945
 | 
     9  | 
  | 
| 
 | 
    10  | 
constdefs
  | 
| 
14963
 | 
    11  | 
  Bij :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) set"
 | 
| 
13945
 | 
    12  | 
    --{*Only extensional functions, since otherwise we get too many.*}
 | 
| 
14963
 | 
    13  | 
  "Bij S \<equiv> extensional S \<inter> {f. bij_betw f S S}"
 | 
| 
13945
 | 
    14  | 
  | 
| 
14963
 | 
    15  | 
  BijGroup :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
 | 
| 
 | 
    16  | 
  "BijGroup S \<equiv>
  | 
| 
 | 
    17  | 
    \<lparr>carrier = Bij S,
  | 
| 
 | 
    18  | 
     mult = \<lambda>g \<in> Bij S. \<lambda>f \<in> Bij S. compose S g f,
  | 
| 
 | 
    19  | 
     one = \<lambda>x \<in> S. x\<rparr>"
  | 
| 
13945
 | 
    20  | 
  | 
| 
 | 
    21  | 
  | 
| 
 | 
    22  | 
declare Id_compose [simp] compose_Id [simp]
  | 
| 
 | 
    23  | 
  | 
| 
14963
 | 
    24  | 
lemma Bij_imp_extensional: "f \<in> Bij S \<Longrightarrow> f \<in> extensional S"
  | 
| 
14666
 | 
    25  | 
  by (simp add: Bij_def)
  | 
| 
13945
 | 
    26  | 
  | 
| 
14963
 | 
    27  | 
lemma Bij_imp_funcset: "f \<in> Bij S \<Longrightarrow> f \<in> S \<rightarrow> S"
  | 
| 
14853
 | 
    28  | 
  by (auto simp add: Bij_def bij_betw_imp_funcset)
  | 
| 
13945
 | 
    29  | 
  | 
| 
 | 
    30  | 
  | 
| 
14666
 | 
    31  | 
subsection {*Bijections Form a Group *}
 | 
| 
13945
 | 
    32  | 
  | 
| 
14963
 | 
    33  | 
lemma restrict_Inv_Bij: "f \<in> Bij S \<Longrightarrow> (\<lambda>x \<in> S. (Inv S f) x) \<in> Bij S"
  | 
| 
14853
 | 
    34  | 
  by (simp add: Bij_def bij_betw_Inv)
  | 
| 
13945
 | 
    35  | 
  | 
| 
 | 
    36  | 
lemma id_Bij: "(\<lambda>x\<in>S. x) \<in> Bij S "
  | 
| 
14853
 | 
    37  | 
  by (auto simp add: Bij_def bij_betw_def inj_on_def)
  | 
| 
13945
 | 
    38  | 
  | 
| 
14963
 | 
    39  | 
lemma compose_Bij: "\<lbrakk>x \<in> Bij S; y \<in> Bij S\<rbrakk> \<Longrightarrow> compose S x y \<in> Bij S"
  | 
| 
14853
 | 
    40  | 
  by (auto simp add: Bij_def bij_betw_compose) 
  | 
| 
13945
 | 
    41  | 
  | 
| 
 | 
    42  | 
lemma Bij_compose_restrict_eq:
  | 
| 
14963
 | 
    43  | 
     "f \<in> Bij S \<Longrightarrow> compose S (restrict (Inv S f) S) f = (\<lambda>x\<in>S. x)"
  | 
| 
14853
 | 
    44  | 
  by (simp add: Bij_def compose_Inv_id)
  | 
| 
13945
 | 
    45  | 
  | 
| 
 | 
    46  | 
theorem group_BijGroup: "group (BijGroup S)"
  | 
| 
14666
 | 
    47  | 
apply (simp add: BijGroup_def)
  | 
| 
13945
 | 
    48  | 
apply (rule groupI)
  | 
| 
 | 
    49  | 
    apply (simp add: compose_Bij)
  | 
| 
 | 
    50  | 
   apply (simp add: id_Bij)
  | 
| 
 | 
    51  | 
  apply (simp add: compose_Bij)
  | 
| 
 | 
    52  | 
  apply (blast intro: compose_assoc [symmetric] Bij_imp_funcset)
  | 
| 
 | 
    53  | 
 apply (simp add: id_Bij Bij_imp_funcset Bij_imp_extensional, simp)
  | 
| 
14666
 | 
    54  | 
apply (blast intro: Bij_compose_restrict_eq restrict_Inv_Bij)
  | 
| 
13945
 | 
    55  | 
done
  | 
| 
 | 
    56  | 
  | 
| 
 | 
    57  | 
  | 
| 
 | 
    58  | 
subsection{*Automorphisms Form a Group*}
 | 
| 
 | 
    59  | 
  | 
| 
14963
 | 
    60  | 
lemma Bij_Inv_mem: "\<lbrakk> f \<in> Bij S;  x \<in> S\<rbrakk> \<Longrightarrow> Inv S f x \<in> S"
  | 
| 
14853
 | 
    61  | 
by (simp add: Bij_def bij_betw_def Inv_mem)
  | 
| 
13945
 | 
    62  | 
  | 
| 
 | 
    63  | 
lemma Bij_Inv_lemma:
  | 
| 
14963
 | 
    64  | 
 assumes eq: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> h(g x y) = g (h x) (h y)"
  | 
| 
 | 
    65  | 
 shows "\<lbrakk>h \<in> Bij S;  g \<in> S \<rightarrow> S \<rightarrow> S;  x \<in> S;  y \<in> S\<rbrakk>
  | 
| 
 | 
    66  | 
        \<Longrightarrow> Inv S h (g x y) = g (Inv S h x) (Inv S h y)"
  | 
| 
14853
 | 
    67  | 
apply (simp add: Bij_def bij_betw_def)
  | 
| 
 | 
    68  | 
apply (subgoal_tac "\<exists>x'\<in>S. \<exists>y'\<in>S. x = h x' & y = h y'", clarify)
  | 
| 
14963
 | 
    69  | 
 apply (simp add: eq [symmetric] Inv_f_f funcset_mem [THEN funcset_mem], blast)
  | 
| 
13945
 | 
    70  | 
done
  | 
| 
 | 
    71  | 
  | 
| 
14963
 | 
    72  | 
  | 
| 
13945
 | 
    73  | 
constdefs
  | 
| 
14963
 | 
    74  | 
  auto :: "('a, 'b) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) set"
 | 
| 
 | 
    75  | 
  "auto G \<equiv> hom G G \<inter> Bij (carrier G)"
  | 
| 
13945
 | 
    76  | 
  | 
| 
14963
 | 
    77  | 
  AutoGroup :: "('a, 'c) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
 | 
| 
 | 
    78  | 
  "AutoGroup G \<equiv> BijGroup (carrier G) \<lparr>carrier := auto G\<rparr>"
  | 
| 
13945
 | 
    79  | 
  | 
| 
14963
 | 
    80  | 
lemma (in group) id_in_auto: "(\<lambda>x \<in> carrier G. x) \<in> auto G"
  | 
| 
14666
 | 
    81  | 
  by (simp add: auto_def hom_def restrictI group.axioms id_Bij)
  | 
| 
13945
 | 
    82  | 
  | 
| 
14963
 | 
    83  | 
lemma (in group) mult_funcset: "mult G \<in> carrier G \<rightarrow> carrier G \<rightarrow> carrier G"
  | 
| 
13945
 | 
    84  | 
  by (simp add:  Pi_I group.axioms)
  | 
| 
 | 
    85  | 
  | 
| 
14963
 | 
    86  | 
lemma (in group) restrict_Inv_hom:
  | 
| 
 | 
    87  | 
      "\<lbrakk>h \<in> hom G G; h \<in> Bij (carrier G)\<rbrakk>
  | 
| 
 | 
    88  | 
       \<Longrightarrow> restrict (Inv (carrier G) h) (carrier G) \<in> hom G G"
  | 
| 
13945
 | 
    89  | 
  by (simp add: hom_def Bij_Inv_mem restrictI mult_funcset
  | 
| 
 | 
    90  | 
                group.axioms Bij_Inv_lemma)
  | 
| 
 | 
    91  | 
  | 
| 
 | 
    92  | 
lemma inv_BijGroup:
  | 
| 
14963
 | 
    93  | 
     "f \<in> Bij S \<Longrightarrow> m_inv (BijGroup S) f = (\<lambda>x \<in> S. (Inv S f) x)"
  | 
| 
13945
 | 
    94  | 
apply (rule group.inv_equality)
  | 
| 
 | 
    95  | 
apply (rule group_BijGroup)
  | 
| 
14666
 | 
    96  | 
apply (simp_all add: BijGroup_def restrict_Inv_Bij Bij_compose_restrict_eq)
  | 
| 
13945
 | 
    97  | 
done
  | 
| 
 | 
    98  | 
  | 
| 
14963
 | 
    99  | 
lemma (in group) subgroup_auto:
  | 
| 
 | 
   100  | 
      "subgroup (auto G) (BijGroup (carrier G))"
  | 
| 
 | 
   101  | 
proof (rule subgroup.intro)
  | 
| 
 | 
   102  | 
  show "auto G \<subseteq> carrier (BijGroup (carrier G))"
  | 
| 
 | 
   103  | 
    by (force simp add: auto_def BijGroup_def)
  | 
| 
 | 
   104  | 
next
  | 
| 
 | 
   105  | 
  fix x y
  | 
| 
 | 
   106  | 
  assume "x \<in> auto G" "y \<in> auto G" 
  | 
| 
 | 
   107  | 
  thus "x \<otimes>\<^bsub>BijGroup (carrier G)\<^esub> y \<in> auto G"
  | 
| 
 | 
   108  | 
    by (force simp add: BijGroup_def is_group auto_def Bij_imp_funcset 
  | 
| 
 | 
   109  | 
                        group.hom_compose compose_Bij)
  | 
| 
 | 
   110  | 
next
  | 
| 
 | 
   111  | 
  show "\<one>\<^bsub>BijGroup (carrier G)\<^esub> \<in> auto G" by (simp add:  BijGroup_def id_in_auto)
  | 
| 
 | 
   112  | 
next
  | 
| 
 | 
   113  | 
  fix x 
  | 
| 
 | 
   114  | 
  assume "x \<in> auto G" 
  | 
| 
 | 
   115  | 
  thus "inv\<^bsub>BijGroup (carrier G)\<^esub> x \<in> auto G"
  | 
| 
 | 
   116  | 
    by (simp del: restrict_apply
  | 
| 
14666
 | 
   117  | 
             add: inv_BijGroup auto_def restrict_Inv_Bij restrict_Inv_hom)
  | 
| 
14963
 | 
   118  | 
qed
  | 
| 
13945
 | 
   119  | 
  | 
| 
14963
 | 
   120  | 
theorem (in group) AutoGroup: "group (AutoGroup G)"
  | 
| 
 | 
   121  | 
by (simp add: AutoGroup_def subgroup.subgroup_is_group subgroup_auto 
  | 
| 
 | 
   122  | 
              group_BijGroup)
  | 
| 
13945
 | 
   123  | 
  | 
| 
 | 
   124  | 
end
  |