src/HOL/Algebra/Bij.thy
author nipkow
Sun, 18 Oct 2009 12:07:25 +0200
changeset 32988 d1d4d7a08a66
parent 31754 b5260f5272a4
child 33057 764547b68538
permissions -rw-r--r--
Inv -> inv_onto, inv abbr. inv_onto UNIV.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
     1
(*  Title:      HOL/Algebra/Bij.thy
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     2
    ID:         $Id$
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     3
    Author:     Florian Kammueller, with new proofs by L C Paulson
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     4
*)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     5
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 16417
diff changeset
     6
theory Bij imports Group begin
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     7
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 16417
diff changeset
     8
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 20318
diff changeset
     9
section {* Bijections of a Set, Permutation and Automorphism Groups *}
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    10
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    11
constdefs
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    12
  Bij :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) set"
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    13
    --{*Only extensional functions, since otherwise we get too many.*}
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    14
  "Bij S \<equiv> extensional S \<inter> {f. bij_betw f S S}"
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    15
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    16
  BijGroup :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    17
  "BijGroup S \<equiv>
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    18
    \<lparr>carrier = Bij S,
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    19
     mult = \<lambda>g \<in> Bij S. \<lambda>f \<in> Bij S. compose S g f,
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    20
     one = \<lambda>x \<in> S. x\<rparr>"
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    21
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    22
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    23
declare Id_compose [simp] compose_Id [simp]
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    24
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    25
lemma Bij_imp_extensional: "f \<in> Bij S \<Longrightarrow> f \<in> extensional S"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 13945
diff changeset
    26
  by (simp add: Bij_def)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    27
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    28
lemma Bij_imp_funcset: "f \<in> Bij S \<Longrightarrow> f \<in> S \<rightarrow> S"
14853
8d710bece29f more on bij_betw
paulson
parents: 14761
diff changeset
    29
  by (auto simp add: Bij_def bij_betw_imp_funcset)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    30
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    31
14666
65f8680c3f16 improved notation;
wenzelm
parents: 13945
diff changeset
    32
subsection {*Bijections Form a Group *}
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    33
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    34
lemma restrict_inv_onto_Bij: "f \<in> Bij S \<Longrightarrow> (\<lambda>x \<in> S. (inv_onto S f) x) \<in> Bij S"
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    35
  by (simp add: Bij_def bij_betw_inv_onto)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    36
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    37
lemma id_Bij: "(\<lambda>x\<in>S. x) \<in> Bij S "
14853
8d710bece29f more on bij_betw
paulson
parents: 14761
diff changeset
    38
  by (auto simp add: Bij_def bij_betw_def inj_on_def)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    39
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    40
lemma compose_Bij: "\<lbrakk>x \<in> Bij S; y \<in> Bij S\<rbrakk> \<Longrightarrow> compose S x y \<in> Bij S"
14853
8d710bece29f more on bij_betw
paulson
parents: 14761
diff changeset
    41
  by (auto simp add: Bij_def bij_betw_compose) 
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    42
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    43
lemma Bij_compose_restrict_eq:
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    44
     "f \<in> Bij S \<Longrightarrow> compose S (restrict (inv_onto S f) S) f = (\<lambda>x\<in>S. x)"
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    45
  by (simp add: Bij_def compose_inv_onto_id)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    46
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    47
theorem group_BijGroup: "group (BijGroup S)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 13945
diff changeset
    48
apply (simp add: BijGroup_def)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    49
apply (rule groupI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    50
    apply (simp add: compose_Bij)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    51
   apply (simp add: id_Bij)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    52
  apply (simp add: compose_Bij)
31754
b5260f5272a4 tuned FuncSet
nipkow
parents: 27717
diff changeset
    53
  apply (blast intro: compose_assoc [symmetric] dest: Bij_imp_funcset)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    54
 apply (simp add: id_Bij Bij_imp_funcset Bij_imp_extensional, simp)
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    55
apply (blast intro: Bij_compose_restrict_eq restrict_inv_onto_Bij)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    56
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    57
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    58
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    59
subsection{*Automorphisms Form a Group*}
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    60
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    61
lemma Bij_inv_onto_mem: "\<lbrakk> f \<in> Bij S;  x \<in> S\<rbrakk> \<Longrightarrow> inv_onto S f x \<in> S"
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    62
by (simp add: Bij_def bij_betw_def inv_onto_into)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    63
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    64
lemma Bij_inv_onto_lemma:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    65
 assumes eq: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> h(g x y) = g (h x) (h y)"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    66
 shows "\<lbrakk>h \<in> Bij S;  g \<in> S \<rightarrow> S \<rightarrow> S;  x \<in> S;  y \<in> S\<rbrakk>
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    67
        \<Longrightarrow> inv_onto S h (g x y) = g (inv_onto S h x) (inv_onto S h y)"
14853
8d710bece29f more on bij_betw
paulson
parents: 14761
diff changeset
    68
apply (simp add: Bij_def bij_betw_def)
8d710bece29f more on bij_betw
paulson
parents: 14761
diff changeset
    69
apply (subgoal_tac "\<exists>x'\<in>S. \<exists>y'\<in>S. x = h x' & y = h y'", clarify)
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    70
 apply (simp add: eq [symmetric] inv_f_f funcset_mem [THEN funcset_mem], blast)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    71
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    72
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    73
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    74
constdefs
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    75
  auto :: "('a, 'b) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) set"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    76
  "auto G \<equiv> hom G G \<inter> Bij (carrier G)"
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    77
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    78
  AutoGroup :: "('a, 'c) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    79
  "AutoGroup G \<equiv> BijGroup (carrier G) \<lparr>carrier := auto G\<rparr>"
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    80
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    81
lemma (in group) id_in_auto: "(\<lambda>x \<in> carrier G. x) \<in> auto G"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 13945
diff changeset
    82
  by (simp add: auto_def hom_def restrictI group.axioms id_Bij)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    83
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    84
lemma (in group) mult_funcset: "mult G \<in> carrier G \<rightarrow> carrier G \<rightarrow> carrier G"
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    85
  by (simp add:  Pi_I group.axioms)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    86
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    87
lemma (in group) restrict_inv_onto_hom:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
    88
      "\<lbrakk>h \<in> hom G G; h \<in> Bij (carrier G)\<rbrakk>
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    89
       \<Longrightarrow> restrict (inv_onto (carrier G) h) (carrier G) \<in> hom G G"
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    90
  by (simp add: hom_def Bij_inv_onto_mem restrictI mult_funcset
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    91
                group.axioms Bij_inv_onto_lemma)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    92
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    93
lemma inv_BijGroup:
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    94
     "f \<in> Bij S \<Longrightarrow> m_inv (BijGroup S) f = (\<lambda>x \<in> S. (inv_onto S f) x)"
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    95
apply (rule group.inv_equality)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    96
apply (rule group_BijGroup)
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
    97
apply (simp_all add:BijGroup_def restrict_inv_onto_Bij Bij_compose_restrict_eq)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    98
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    99
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   100
lemma (in group) subgroup_auto:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   101
      "subgroup (auto G) (BijGroup (carrier G))"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   102
proof (rule subgroup.intro)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   103
  show "auto G \<subseteq> carrier (BijGroup (carrier G))"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   104
    by (force simp add: auto_def BijGroup_def)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   105
next
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   106
  fix x y
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   107
  assume "x \<in> auto G" "y \<in> auto G" 
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   108
  thus "x \<otimes>\<^bsub>BijGroup (carrier G)\<^esub> y \<in> auto G"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   109
    by (force simp add: BijGroup_def is_group auto_def Bij_imp_funcset 
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   110
                        group.hom_compose compose_Bij)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   111
next
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   112
  show "\<one>\<^bsub>BijGroup (carrier G)\<^esub> \<in> auto G" by (simp add:  BijGroup_def id_in_auto)
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   113
next
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   114
  fix x 
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   115
  assume "x \<in> auto G" 
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   116
  thus "inv\<^bsub>BijGroup (carrier G)\<^esub> x \<in> auto G"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   117
    by (simp del: restrict_apply
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 31754
diff changeset
   118
        add: inv_BijGroup auto_def restrict_inv_onto_Bij restrict_inv_onto_hom)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   119
qed
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   120
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   121
theorem (in group) AutoGroup: "group (AutoGroup G)"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   122
by (simp add: AutoGroup_def subgroup.subgroup_is_group subgroup_auto 
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14853
diff changeset
   123
              group_BijGroup)
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   124
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   125
end