| author | paulson | 
| Mon, 04 Feb 2002 13:16:54 +0100 | |
| changeset 12867 | 5c900a821a7c | 
| parent 12459 | 6978ab7cac64 | 
| permissions | -rw-r--r-- | 
| 11394 | 1 | (* Title: HOL/GroupTheory/DirProd | 
| 2 | ID: $Id$ | |
| 3 | Author: Florian Kammueller, with new proofs by L C Paulson | |
| 4 | Copyright 1998-2001 University of Cambridge | |
| 5 | ||
| 6 | Direct product of two groups | |
| 7 | *) | |
| 8 | ||
| 9 | Open_locale "prodgroup"; | |
| 10 | ||
| 11 | val e'_def = thm "e'_def"; | |
| 12 | val binop'_def = thm "binop'_def"; | |
| 13 | val inv'_def = thm "inv'_def"; | |
| 14 | val P_def = thm "P_def"; | |
| 15 | val Group_G' = thm "Group_G'"; | |
| 16 | ||
| 17 | ||
| 18 | Addsimps [P_def, Group_G', Group_G]; | |
| 19 | ||
| 20 | Goal "(P.<cr>) = ((G.<cr>) \\<times> (G'.<cr>))"; | |
| 21 | by (afs [ProdGroup_def] 1); | |
| 22 | qed "P_carrier"; | |
| 23 | ||
| 24 | Goal "(P.<f>) = \ | |
| 12459 
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
 wenzelm parents: 
11443diff
changeset | 25 | \ (%(x, x'): (P.<cr>). %(y, y'): (P.<cr>). ( x ## y, x' ##' y'))"; | 
| 11394 | 26 | by (afs [ProdGroup_def, binop_def, binop'_def] 1); | 
| 27 | qed "P_bin_op"; | |
| 28 | ||
| 12459 
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
 wenzelm parents: 
11443diff
changeset | 29 | Goal "(P.<inv>) = (%(x, y): (P.<cr>). (i x, i' y))"; | 
| 11394 | 30 | by (afs [ProdGroup_def, inv_def, inv'_def] 1); | 
| 31 | qed "P_inverse"; | |
| 32 | ||
| 33 | Goal "(P.<e>) = (G.<e>, G'.<e>)"; | |
| 34 | by (afs [ProdGroup_def, e_def, e'_def] 1); | |
| 35 | qed "P_unit"; | |
| 36 | ||
| 37 | Goal "P = (| carrier = P.<cr>, \ | |
| 12459 
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
 wenzelm parents: 
11443diff
changeset | 38 | \ bin_op = (%(x, x'): (P.<cr>). %(y, y'): (P.<cr>).\ | 
| 11394 | 39 | \ (x ## y, x' ##' y')), \ | 
| 12459 
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
 wenzelm parents: 
11443diff
changeset | 40 | \ inverse = (%(x, y): (P.<cr>). (i x, i' y)), \ | 
| 11394 | 41 | \ unit = P.<e> |)"; | 
| 42 | by (afs [ProdGroup_def, P_carrier, P_bin_op, P_inverse, P_unit] 1); | |
| 43 | by (afs [binop_def, binop'_def, inv_def, inv'_def] 1); | |
| 44 | qed "P_defI"; | |
| 45 | val P_DefI = export P_defI; | |
| 46 | ||
| 47 | Delsimps [P_def]; | |
| 48 | ||
| 49 | Goal "(P.<f>) : (P.<cr>) \\<rightarrow> (P.<cr>) \\<rightarrow> (P.<cr>)"; | |
| 50 | by (auto_tac (claset() addSIs [restrictI], | |
| 51 | simpset() addsimps [P_bin_op, P_carrier, binop'_def, | |
| 52 | bin_op_closed])); | |
| 53 | qed "bin_op_prod_closed"; | |
| 54 | ||
| 55 | Goal "(P.<inv>) : (P.<cr>) \\<rightarrow> (P.<cr>)"; | |
| 56 | by (auto_tac (claset() addSIs [restrictI], | |
| 57 | simpset() addsimps [P_inverse, P_carrier, inv_closed, | |
| 58 | inv'_def, inverse_closed])); | |
| 59 | qed "inverse_prod_closed"; | |
| 60 | ||
| 61 | (* MAIN PROOF *) | |
| 62 | Goal "P : Group"; | |
| 63 | by (stac P_defI 1); | |
| 64 | by (rtac GroupI 1); | |
| 65 | by (auto_tac (claset(), | |
| 66 | simpset() addsimps ([P_carrier,P_bin_op,P_inverse,P_unit] RL [sym]))); | |
| 67 | (* 1. *) | |
| 68 | by (rtac bin_op_prod_closed 1); | |
| 69 | (* 2. *) | |
| 70 | by (rtac inverse_prod_closed 1); | |
| 71 | (* 3. *) | |
| 72 | by (afs [P_carrier, P_unit, export e_closed] 1); | |
| 73 | (* 4. *) | |
| 74 | by (afs [P_carrier, P_bin_op, P_inverse, P_unit, Group_G' RS inverse_closed, | |
| 75 | inv'_def, e_def, binop'_def, Group_G' RS (export inv_ax2)] 1); | |
| 76 | (* 5 *) | |
| 77 | by (afs [P_carrier,P_bin_op,P_unit, Group_G' RS unit_closed, export e_ax1, | |
| 78 | binop_def, binop'_def] 1); | |
| 79 | (* 6 *) | |
| 80 | by (afs [P_carrier,P_bin_op, Group_G' RS bin_op_closed, | |
| 81 | binop'_def, binop_assoc,export binop_assoc] 1); | |
| 82 | qed "prodgroup_is_group"; | |
| 83 | val ProdGroup_is_Group = export prodgroup_is_group; | |
| 84 | ||
| 11443 | 85 | Delsimps [P_def, Group_G', Group_G]; | 
| 11394 | 86 | |
| 87 | Close_locale "prodgroup"; | |
| 88 | Close_locale "r_group"; | |
| 89 | Close_locale "group"; | |
| 90 | ||
| 91 | Goal "ProdGroup : Group \\<rightarrow> Group \\<rightarrow> Group"; | |
| 92 | by (REPEAT (ares_tac [funcsetI, ProdGroup_is_Group] 1)); | |
| 93 | by (auto_tac (claset(), simpset() addsimps [ProdGroup_def])); | |
| 94 | qed "ProdGroup_arity"; |