author | nipkow |
Mon, 13 May 2002 15:27:28 +0200 | |
changeset 13145 | 59bc43b51aa2 |
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 |
DirProd = Coset + |
|
10 |
||
11 |
consts |
|
12 |
ProdGroup :: "(['a grouptype, 'b grouptype] => ('a * 'b) grouptype)" |
|
13 |
||
14 |
defs |
|
12459
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
wenzelm
parents:
11394
diff
changeset
|
15 |
ProdGroup_def "ProdGroup == %G: Group. %G': Group. |
11394 | 16 |
(| carrier = ((G.<cr>) \\<times> (G'.<cr>)), |
12459
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
wenzelm
parents:
11394
diff
changeset
|
17 |
bin_op = (%(x, x'): ((G.<cr>) \\<times> (G'.<cr>)). |
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
wenzelm
parents:
11394
diff
changeset
|
18 |
%(y, y'): ((G.<cr>) \\<times> (G'.<cr>)). |
11394 | 19 |
((G.<f>) x y,(G'.<f>) x' y')), |
12459
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
wenzelm
parents:
11394
diff
changeset
|
20 |
inverse = (%(x, y): ((G.<cr>) \\<times> (G'.<cr>)). ((G.<inv>) x, (G'.<inv>) y)), |
11394 | 21 |
unit = ((G.<e>),(G'.<e>)) |)" |
22 |
||
23 |
syntax |
|
24 |
"@Pro" :: "['a grouptype, 'b grouptype] => ('a * 'b) grouptype" ("<|_,_|>" [60,61]60) |
|
25 |
||
26 |
translations |
|
27 |
"<| G , G' |>" == "ProdGroup G G'" |
|
28 |
||
29 |
locale r_group = group + |
|
30 |
fixes |
|
31 |
G' :: "'b grouptype" |
|
32 |
e' :: "'b" |
|
33 |
binop' :: "'b => 'b => 'b" ("(_ ##' _)" [80,81]80 ) |
|
34 |
INV' :: "'b => 'b" ("i' (_)" [90]91) |
|
35 |
assumes |
|
36 |
Group_G' "G' : Group" |
|
37 |
defines |
|
38 |
e'_def "e' == unit G'" |
|
39 |
binop'_def "x ##' y == bin_op G' x y" |
|
40 |
inv'_def "i'(x) == inverse G' x" |
|
41 |
||
42 |
locale prodgroup = r_group + |
|
43 |
fixes |
|
44 |
P :: "('a * 'b) grouptype" |
|
45 |
defines |
|
46 |
P_def "P == <| G, G' |>" |
|
47 |
end |