author | nipkow |
Mon, 13 May 2002 15:27:28 +0200 | |
changeset 13145 | 59bc43b51aa2 |
parent 12459 | 6978ab7cac64 |
child 13583 | 5fcc8bf538ee |
permissions | -rw-r--r-- |
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11448
diff
changeset
|
1 |
(* Title: HOL/GroupTheory/Bij |
11448 | 2 |
ID: $Id$ |
3 |
Author: Florian Kammueller, with new proofs by L C Paulson |
|
4 |
Copyright 1998-2001 University of Cambridge |
|
5 |
||
6 |
Bijections of a set and the group of bijections |
|
7 |
Sigma version with locales |
|
8 |
*) |
|
9 |
||
10 |
Bij = Group + |
|
11 |
||
12 |
constdefs |
|
13 |
Bij :: "'a set => (('a => 'a)set)" |
|
14 |
"Bij S == {f. f \\<in> S \\<rightarrow> S & f`S = S & inj_on f S}" |
|
15 |
||
16 |
constdefs |
|
17 |
BijGroup :: "'a set => (('a => 'a) grouptype)" |
|
18 |
"BijGroup S == (| carrier = Bij S, |
|
12459
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
wenzelm
parents:
11451
diff
changeset
|
19 |
bin_op = %g: Bij S. %f: Bij S. compose S g f, |
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
wenzelm
parents:
11451
diff
changeset
|
20 |
inverse = %f: Bij S. %x: S. (Inv S f) x, |
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
wenzelm
parents:
11451
diff
changeset
|
21 |
unit = %x: S. x |)" |
11448 | 22 |
|
23 |
locale bij = |
|
24 |
fixes |
|
25 |
S :: "'a set" |
|
26 |
B :: "('a => 'a)set" |
|
27 |
comp :: "[('a => 'a),('a => 'a)]=>('a => 'a)" (infixr "o''" 80) |
|
28 |
inv' :: "('a => 'a)=>('a => 'a)" |
|
29 |
e' :: "('a => 'a)" |
|
30 |
defines |
|
31 |
B_def "B == Bij S" |
|
32 |
o'_def "g o' f == compose S g f" |
|
33 |
inv'_def "inv' f == Inv S f" |
|
12459
6978ab7cac64
bounded abstraction now uses syntax "%" / "\<lambda>" instead of "lam";
wenzelm
parents:
11451
diff
changeset
|
34 |
e'_def "e' == (%x: S. x)" |
11448 | 35 |
|
36 |
locale bijgroup = bij + |
|
37 |
fixes |
|
38 |
BG :: "('a => 'a) grouptype" |
|
39 |
defines |
|
40 |
BG_def "BG == BijGroup S" |
|
41 |
end |
|
42 |