src/HOL/Algebra/Bij.thy
author ballarin
Fri, 02 May 2003 20:02:50 +0200
changeset 13949 0ce528cd6f19
parent 13945 5433b2755e98
child 14666 65f8680c3f16
permissions -rw-r--r--
HOL-Algebra complete for release Isabelle2003 (modulo section headers).
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13945
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Algebra/Bij
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
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     6
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     7
header{*Bijections of a Set, Permutation Groups, Automorphism Groups*} 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     8
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
     9
theory Bij = Group:
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
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    12
  Bij :: "'a set => (('a => 'a)set)" 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    13
    --{*Only extensional functions, since otherwise we get too many.*}
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    14
    "Bij S == extensional S \<inter> {f. f`S = S & inj_on f S}"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    15
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    16
   BijGroup ::  "'a set => (('a => 'a) monoid)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    17
   "BijGroup S == (| carrier = Bij S, 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    18
		     mult  = %g: Bij S. %f: Bij S. compose S g f,
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    19
		     one = %x: S. x |)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    20
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    21
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    22
declare Id_compose [simp] compose_Id [simp]
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    23
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    24
lemma Bij_imp_extensional: "f \<in> Bij S ==> f \<in> extensional S"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    25
by (simp add: Bij_def)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    26
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    27
lemma Bij_imp_funcset: "f \<in> Bij S ==> f \<in> S -> S"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    28
by (auto simp add: Bij_def Pi_def)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    29
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    30
lemma Bij_imp_apply: "f \<in> Bij S ==> f ` S = S"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    31
by (simp add: Bij_def)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    32
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    33
lemma Bij_imp_inj_on: "f \<in> Bij S ==> inj_on f S"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    34
by (simp add: Bij_def)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    35
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    36
lemma BijI: "[| f \<in> extensional(S); f`S = S; inj_on f S |] ==> f \<in> Bij S"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    37
by (simp add: Bij_def)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    38
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    39
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    40
subsection{*Bijections Form a Group*}
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    41
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    42
lemma restrict_Inv_Bij: "f \<in> Bij S ==> (%x:S. (Inv S f) x) \<in> Bij S"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    43
apply (simp add: Bij_def)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    44
apply (intro conjI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    45
txt{*Proving @{term "restrict (Inv S f) S ` S = S"}*}
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    46
 apply (rule equalityI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    47
  apply (force simp add: Inv_mem) --{*first inclusion*}
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    48
 apply (rule subsetI)   --{*second inclusion*}
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    49
 apply (rule_tac x = "f x" in image_eqI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    50
  apply (force intro:  simp add: Inv_f_f, blast)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    51
txt{*Remaining goal: @{term "inj_on (restrict (Inv S f) S) S"}*}
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    52
apply (rule inj_onI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    53
apply (auto elim: Inv_injective)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    54
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    55
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    56
lemma id_Bij: "(\<lambda>x\<in>S. x) \<in> Bij S "
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    57
apply (rule BijI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    58
apply (auto simp add: inj_on_def)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    59
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    60
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    61
lemma compose_Bij: "[| x \<in> Bij S; y \<in> Bij S|] ==> compose S x y \<in> Bij S"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    62
apply (rule BijI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    63
  apply (simp add: compose_extensional) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    64
 apply (blast del: equalityI
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    65
              intro: surj_compose dest: Bij_imp_apply Bij_imp_inj_on)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    66
apply (blast intro: inj_on_compose dest: Bij_imp_apply Bij_imp_inj_on)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    67
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    68
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    69
lemma Bij_compose_restrict_eq:
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    70
     "f \<in> Bij S ==> compose S (restrict (Inv S f) S) f = (\<lambda>x\<in>S. x)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    71
apply (rule compose_Inv_id)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    72
 apply (simp add: Bij_imp_inj_on)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    73
apply (simp add: Bij_imp_apply) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    74
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    75
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    76
theorem group_BijGroup: "group (BijGroup S)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    77
apply (simp add: BijGroup_def) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    78
apply (rule groupI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    79
    apply (simp add: compose_Bij)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    80
   apply (simp add: id_Bij)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    81
  apply (simp add: compose_Bij)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    82
  apply (blast intro: compose_assoc [symmetric] Bij_imp_funcset)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    83
 apply (simp add: id_Bij Bij_imp_funcset Bij_imp_extensional, simp)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    84
apply (blast intro: Bij_compose_restrict_eq restrict_Inv_Bij) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    85
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    86
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    87
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    88
subsection{*Automorphisms Form a Group*}
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    89
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    90
lemma Bij_Inv_mem: "[|  f \<in> Bij S;  x : S |] ==> Inv S f x : S"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    91
by (simp add: Bij_def Inv_mem) 
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 Bij_Inv_lemma:
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    94
 assumes eq: "!!x y. [|x \<in> S; y \<in> S|] ==> h(g x y) = g (h x) (h y)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    95
 shows "[| h \<in> Bij S;  g \<in> S \<rightarrow> S \<rightarrow> S;  x \<in> S;  y \<in> S |]        
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    96
        ==> Inv S h (g x y) = g (Inv S h x) (Inv S h y)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    97
apply (simp add: Bij_def) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    98
apply (subgoal_tac "EX x':S. EX y':S. x = h x' & y = h y'", clarify)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
    99
 apply (simp add: eq [symmetric] Inv_f_f funcset_mem [THEN funcset_mem], blast)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   100
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   101
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   102
constdefs
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   103
 auto :: "('a,'b) monoid_scheme => ('a => 'a)set"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   104
  "auto G == hom G G \<inter> Bij (carrier G)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   105
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   106
  AutoGroup :: "[('a,'c) monoid_scheme] => ('a=>'a) monoid"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   107
  "AutoGroup G == BijGroup (carrier G) (|carrier := auto G |)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   108
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   109
lemma id_in_auto: "group G ==> (%x: carrier G. x) \<in> auto G"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   110
  by (simp add: auto_def hom_def restrictI group.axioms id_Bij) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   111
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   112
lemma mult_funcset: "group G ==> mult G \<in> carrier G -> carrier G -> carrier G"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   113
  by (simp add:  Pi_I group.axioms)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   114
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   115
lemma restrict_Inv_hom:
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   116
      "[|group G; h \<in> hom G G; h \<in> Bij (carrier G)|]
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   117
       ==> restrict (Inv (carrier G) h) (carrier G) \<in> hom G G"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   118
  by (simp add: hom_def Bij_Inv_mem restrictI mult_funcset
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   119
                group.axioms Bij_Inv_lemma)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   120
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   121
lemma inv_BijGroup:
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   122
     "f \<in> Bij S ==> m_inv (BijGroup S) f = (%x: S. (Inv S f) x)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   123
apply (rule group.inv_equality)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   124
apply (rule group_BijGroup)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   125
apply (simp_all add: BijGroup_def restrict_Inv_Bij Bij_compose_restrict_eq)  
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   126
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   127
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   128
lemma subgroup_auto:
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   129
      "group G ==> subgroup (auto G) (BijGroup (carrier G))"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   130
apply (rule group.subgroupI) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   131
    apply (rule group_BijGroup) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   132
   apply (force simp add: auto_def BijGroup_def) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   133
  apply (blast intro: dest: id_in_auto) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   134
 apply (simp del: restrict_apply
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   135
	     add: inv_BijGroup auto_def restrict_Inv_Bij restrict_Inv_hom) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   136
apply (simp add: BijGroup_def auto_def Bij_imp_funcset compose_hom compose_Bij)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   137
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   138
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   139
theorem AutoGroup: "group G ==> group (AutoGroup G)"
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   140
apply (simp add: AutoGroup_def) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   141
apply (rule Group.subgroup.groupI)
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   142
apply (erule subgroup_auto)  
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   143
apply (insert Bij.group_BijGroup [of "carrier G"]) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   144
apply (simp_all add: group_def) 
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   145
done
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   146
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   147
end
5433b2755e98 moved Bij.thy from HOL/GroupTheory
paulson
parents:
diff changeset
   148