src/HOL/GroupTheory/DirProd.thy
author nipkow
Mon, 13 May 2002 15:27:28 +0200
changeset 13145 59bc43b51aa2
parent 12459 6978ab7cac64
permissions -rw-r--r--
*** empty log message ***
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11394
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     1
(*  Title:      HOL/GroupTheory/DirProd
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     2
    ID:         $Id$
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     3
    Author:     Florian Kammueller, with new proofs by L C Paulson
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     4
    Copyright   1998-2001  University of Cambridge
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     5
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     6
Direct product of two groups
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     7
*)
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     8
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
     9
DirProd = Coset +
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    10
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    11
consts
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    12
  ProdGroup :: "(['a grouptype, 'b grouptype] => ('a * 'b) grouptype)"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    13
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    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
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    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
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    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
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    21
       unit = ((G.<e>),(G'.<e>)) |)"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    22
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    23
syntax
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    24
  "@Pro" :: "['a grouptype, 'b grouptype] => ('a * 'b) grouptype" ("<|_,_|>" [60,61]60)
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    25
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    26
translations
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    27
  "<| G , G' |>" == "ProdGroup G G'"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    28
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    29
locale r_group = group +
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    30
  fixes
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    31
    G'        :: "'b grouptype"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    32
    e'        :: "'b"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    33
    binop'    :: "'b => 'b => 'b" 	("(_ ##' _)" [80,81]80 )
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    34
    INV'      :: "'b => 'b"              ("i' (_)"      [90]91)
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    35
  assumes 
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    36
    Group_G' "G' : Group"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    37
  defines
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    38
    e'_def  "e' == unit G'"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    39
    binop'_def "x ##' y == bin_op G' x y"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    40
    inv'_def "i'(x) == inverse G' x"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    41
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    42
locale prodgroup = r_group +
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    43
  fixes 
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    44
    P :: "('a * 'b) grouptype"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    45
  defines 
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    46
    P_def "P == <| G, G' |>"
e88c2c89f98e Locale-based group theory proofs
paulson
parents:
diff changeset
    47
end