moving Bij.thy from GroupTheory to Algebra
authorpaulson
Thu, 01 May 2003 10:29:44 +0200
changeset 13943 83d842ccd4aa
parent 13942 dc93e3a68142
child 13944 9b34607cd83e
moving Bij.thy from GroupTheory to Algebra
src/HOL/Algebra/Coset.thy
src/HOL/Algebra/Group.thy
src/HOL/Algebra/ROOT.ML
src/HOL/GroupTheory/Bij.thy
src/HOL/GroupTheory/ROOT.ML
src/HOL/IsaMakefile
--- a/src/HOL/Algebra/Coset.thy	Thu May 01 08:39:37 2003 +0200
+++ b/src/HOL/Algebra/Coset.thy	Thu May 01 10:29:44 2003 +0200
@@ -491,23 +491,4 @@
 apply (auto dest: coset.setrcos_inv_mult_group_eq simp add: setinv_closed)
 done
 
-(*????????????????
-theorem factorgroup_is_group: "H <| G ==> group (G Mod H)"
-apply (frule normal_imp_coset) 
-apply (simp add: FactGroup_def) 
-apply (rule group.intro)
-apply (rule magma.intro) 
-apply (simp add: ); 
-  apply (simp add: restrictI coset.setmult_closed) 
- apply (rule semigroup.intro) 
-  apply (simp add: restrictI coset.setmult_closed) 
- apply (simp add: coset.setmult_closed coset.setrcos_assoc)  
-apply (rule group_axioms.intro) 
-   apply (simp add: restrictI setinv_closed) 
-  apply (simp add: normal_imp_subgroup subgroup_in_rcosets) 
- apply (simp add: setinv_closed coset.setrcos_inv_mult_group_eq) 
-apply (simp add: normal_imp_subgroup subgroup_in_rcosets coset.setrcos_mult_eq)
-done
-*)
-
 end
--- a/src/HOL/Algebra/Group.thy	Thu May 01 08:39:37 2003 +0200
+++ b/src/HOL/Algebra/Group.thy	Thu May 01 10:29:44 2003 +0200
@@ -99,29 +99,23 @@
 
 lemma (in monoid) Units_inv_closed [intro, simp]:
   "x \<in> Units G ==> inv x \<in> carrier G"
-  apply (unfold Units_def m_inv_def)
-  apply auto
+  apply (unfold Units_def m_inv_def, auto)
   apply (rule theI2, fast)
-   apply (fast intro: inv_unique)
-  apply fast
+   apply (fast intro: inv_unique, fast)
   done
 
 lemma (in monoid) Units_l_inv:
   "x \<in> Units G ==> inv x \<otimes> x = \<one>"
-  apply (unfold Units_def m_inv_def)
-  apply auto
+  apply (unfold Units_def m_inv_def, auto)
   apply (rule theI2, fast)
-   apply (fast intro: inv_unique)
-  apply fast
+   apply (fast intro: inv_unique, fast)
   done
 
 lemma (in monoid) Units_r_inv:
   "x \<in> Units G ==> x \<otimes> inv x = \<one>"
-  apply (unfold Units_def m_inv_def)
-  apply auto
+  apply (unfold Units_def m_inv_def, auto)
   apply (rule theI2, fast)
-   apply (fast intro: inv_unique)
-  apply fast
+   apply (fast intro: inv_unique, fast)
   done
 
 lemma (in monoid) Units_inv_Units [intro, simp]:
@@ -354,6 +348,14 @@
   "[| x \<otimes> y = \<one>; x \<in> carrier G; y \<in> carrier G |] ==> y \<otimes> x = \<one>"
   by (rule Units_inv_comm) auto                          
 
+lemma (in group) m_inv_equality:
+     "[|y \<otimes> x = \<one>; x \<in> carrier G; y \<in> carrier G|] ==> inv x = y"
+apply (simp add: m_inv_def)
+apply (rule the_equality)
+ apply (simp add: inv_comm [of y x]) 
+apply (rule r_cancel [THEN iffD1], auto) 
+done
+
 text {* Power *}
 
 lemma (in group) int_pow_def2:
@@ -594,6 +596,15 @@
   "[| h \<in> hom G H; x \<in> carrier G |] ==> h x \<in> carrier H"
   by (auto simp add: hom_def funcset_mem)
 
+lemma compose_hom:
+     "[|group G; h \<in> hom G G; h' \<in> hom G G; h' \<in> carrier G -> carrier G|]
+      ==> compose (carrier G) h h' \<in> hom G G"
+apply (simp (no_asm_simp) add: hom_def)
+apply (intro conjI) 
+ apply (force simp add: funcset_compose hom_def)
+apply (simp add: compose_def group.axioms hom_mult funcset_mem) 
+done
+
 locale group_hom = group G + group H + var h +
   assumes homh: "h \<in> hom G H"
   notes hom_mult [simp] = hom_mult [OF homh]
--- a/src/HOL/Algebra/ROOT.ML	Thu May 01 08:39:37 2003 +0200
+++ b/src/HOL/Algebra/ROOT.ML	Thu May 01 10:29:44 2003 +0200
@@ -8,6 +8,7 @@
 
 no_document use_thy "FuncSet";
 use_thy "Sylow";		(* Groups *)
+use_thy "Bij";			(* Automorphism Groups *)
 use_thy "UnivPoly";		(* Rings and polynomials *)
 
 (* Old development, based on axiomatic type classes.
--- a/src/HOL/GroupTheory/Bij.thy	Thu May 01 08:39:37 2003 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,131 +0,0 @@
-(*  Title:      HOL/GroupTheory/Bij
-    ID:         $Id$
-    Author:     Florian Kammueller, with new proofs by L C Paulson
-*)
-
-
-header{*Bijections of a Set, Permutation Groups, Automorphism Groups*} 
-
-theory Bij = Group:
-
-constdefs
-  Bij :: "'a set => (('a => 'a)set)" 
-    --{*Only extensional functions, since otherwise we get too many.*}
-    "Bij S == extensional S \<inter> {f. f`S = S & inj_on f S}"
-
-   BijGroup ::  "'a set => (('a => 'a) group)"
-   "BijGroup S == (| carrier = Bij S, 
-		     sum  = %g: Bij S. %f: Bij S. compose S g f,
-		     gminus = %f: Bij S. %x: S. (Inv S f) x, 
-		     zero = %x: S. x |)"
-
-
-declare Id_compose [simp] compose_Id [simp]
-
-lemma Bij_imp_extensional: "f \<in> Bij S ==> f \<in> extensional S"
-by (simp add: Bij_def)
-
-lemma Bij_imp_funcset: "f \<in> Bij S ==> f \<in> S -> S"
-by (auto simp add: Bij_def Pi_def)
-
-lemma Bij_imp_apply: "f \<in> Bij S ==> f ` S = S"
-by (simp add: Bij_def)
-
-lemma Bij_imp_inj_on: "f \<in> Bij S ==> inj_on f S"
-by (simp add: Bij_def)
-
-lemma BijI: "[| f \<in> extensional(S); f`S = S; inj_on f S |] ==> f \<in> Bij S"
-by (simp add: Bij_def)
-
-
-subsection{*Bijections Form a Group*}
-
-lemma restrict_Inv_Bij: "f \<in> Bij S ==> (%x:S. (Inv S f) x) \<in> Bij S"
-apply (simp add: Bij_def)
-apply (intro conjI)
-txt{*Proving @{term "restrict (Inv S f) S ` S = S"}*}
- apply (rule equalityI)
-  apply (force simp add: Inv_mem) --{*first inclusion*}
- apply (rule subsetI)   --{*second inclusion*}
- apply (rule_tac x = "f x" in image_eqI)
-  apply (force intro:  simp add: Inv_f_f)
- apply blast
-txt{*Remaining goal: @{term "inj_on (restrict (Inv S f) S) S"}*}
-apply (rule inj_onI)
-apply (auto elim: Inv_injective)
-done
-
-lemma id_Bij: "(\<lambda>x\<in>S. x) \<in> Bij S "
-apply (rule BijI)
-apply (auto simp add: inj_on_def)
-done
-
-lemma compose_Bij: "[| x \<in> Bij S; y \<in> Bij S|] ==> compose S x y \<in> Bij S"
-apply (rule BijI)
-  apply (simp add: compose_extensional) 
- apply (blast del: equalityI
-              intro: surj_compose dest: Bij_imp_apply Bij_imp_inj_on)
-apply (blast intro: inj_on_compose dest: Bij_imp_apply Bij_imp_inj_on)
-done
-
-theorem group_BijGroup: "group (BijGroup S)"
-apply (simp add: group_def semigroup_def group_axioms_def 
-                 BijGroup_def restrictI compose_Bij restrict_Inv_Bij id_Bij)
-apply (auto intro!: compose_Bij) 
-  apply (blast intro: compose_assoc [symmetric] Bij_imp_funcset)
- apply (simp add: Bij_def compose_Inv_id) 
-apply (simp add: Id_compose Bij_imp_funcset Bij_imp_extensional) 
-done
-
-
-subsection{*Automorphisms Form a Group*}
-
-lemma Bij_Inv_mem: "[|  f \<in> Bij S;  x : S |] ==> Inv S f x : S"
-by (simp add: Bij_def Inv_mem) 
-
-lemma Bij_Inv_lemma:
- assumes eq: "!!x y. [|x \<in> S; y \<in> S|] ==> h(g x y) = g (h x) (h y)"
- shows "[| h \<in> Bij S;  g \<in> S \<rightarrow> S \<rightarrow> S;  x \<in> S;  y \<in> S |]        
-        ==> Inv S h (g x y) = g (Inv S h x) (Inv S h y)"
-apply (simp add: Bij_def) 
-apply (subgoal_tac "EX x':S. EX y':S. x = h x' & y = h y'")
- apply clarify
- apply (simp add: eq [symmetric] Inv_f_f funcset_mem [THEN funcset_mem], blast) 
-done
-
-constdefs
- auto :: "('a,'b)group_scheme => ('a => 'a)set"
-  "auto G == hom G G Int Bij (carrier G)"
-
-  AutoGroup :: "[('a,'c) group_scheme] => ('a=>'a) group"
-  "AutoGroup G == BijGroup (carrier G) (|carrier := auto G |)"
-
-lemma id_in_auto: "group G ==> (%x: carrier G. x) \<in> auto G"
-by (simp add: auto_def hom_def restrictI semigroup.sum_closed 
-              group.axioms id_Bij) 
-
-lemma restrict_Inv_hom:
-      "[|group G; h \<in> hom G G; h \<in> Bij (carrier G)|]
-       ==> restrict (Inv (carrier G) h) (carrier G) \<in> hom G G"
-by (simp add: hom_def Bij_Inv_mem restrictI semigroup.sum_closed 
-              semigroup.sum_funcset group.axioms Bij_Inv_lemma)
-
-lemma subgroup_auto:
-      "group G ==> subgroup (auto G) (BijGroup (carrier G))"
-apply (rule group.subgroupI) 
-    apply (rule group_BijGroup) 
-   apply (force simp add: auto_def BijGroup_def) 
-  apply (blast intro: dest: id_in_auto) 
- apply (simp add: auto_def BijGroup_def restrict_Inv_Bij
-                  restrict_Inv_hom) 
-apply (simp add: auto_def BijGroup_def compose_Bij)
-apply (simp add: hom_def compose_def Pi_def group.axioms)
-done
-
-theorem AutoGroup: "group G ==> group (AutoGroup G)"
-apply (drule subgroup_auto) 
-apply (simp add: subgroup_def AutoGroup_def) 
-done
-
-end
-
--- a/src/HOL/GroupTheory/ROOT.ML	Thu May 01 08:39:37 2003 +0200
+++ b/src/HOL/GroupTheory/ROOT.ML	Thu May 01 10:29:44 2003 +0200
@@ -1,3 +1,3 @@
 no_document use_thy "FuncSet";
 
-use_thy "Bij";
+use_thy "Group";
--- a/src/HOL/IsaMakefile	Thu May 01 08:39:37 2003 +0200
+++ b/src/HOL/IsaMakefile	Thu May 01 10:29:44 2003 +0200
@@ -284,7 +284,6 @@
 
 $(LOG)/HOL-GroupTheory.gz: $(OUT)/HOL \
   Library/Primes.thy Library/FuncSet.thy \
-  GroupTheory/Bij.thy \
   GroupTheory/Group.thy \
   GroupTheory/ROOT.ML \
   GroupTheory/document/root.tex
@@ -341,6 +340,7 @@
 HOL-Algebra: HOL $(LOG)/HOL-Algebra.gz
 
 $(LOG)/HOL-Algebra.gz: $(OUT)/HOL Algebra/ROOT.ML \
+  Algebra/Bij.thy \
   Algebra/CRing.thy \
   Algebra/Coset.thy \
   Algebra/Exponent.thy \