| author | wenzelm | 
| Tue, 07 Aug 2001 21:27:00 +0200 | |
| changeset 11472 | d08d4e17a5f6 | 
| parent 11451 | 8abfb4f7bd02 | 
| child 12459 | 6978ab7cac64 | 
| permissions | -rw-r--r-- | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11448diff
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, | |
| 19 | bin_op = lam g: Bij S. lam f: Bij S. compose S g f, | |
| 20 | inverse = lam f: Bij S. lam x: S. (Inv S f) x, | |
| 21 | unit = lam x: S. x |)" | |
| 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" | |
| 34 | e'_def "e' == (lam x: S. x)" | |
| 35 | ||
| 36 | locale bijgroup = bij + | |
| 37 | fixes | |
| 38 |     BG :: "('a => 'a) grouptype"
 | |
| 39 | defines | |
| 40 | BG_def "BG == BijGroup S" | |
| 41 | end | |
| 42 |