| author | wenzelm | 
| Thu, 14 Jul 2005 19:28:16 +0200 | |
| changeset 16834 | 71d87aeebb57 | 
| parent 13735 | 7de9342aca7a | 
| child 17479 | 68a7acb5f22e | 
| permissions | -rw-r--r-- | 
| 7998 | 1 | (* | 
| 2 | Factorisation within a factorial domain | |
| 3 | $Id$ | |
| 4 | Author: Clemens Ballarin, started 25 November 1997 | |
| 5 | *) | |
| 6 | ||
| 7 | Factor = Ring + | |
| 8 | ||
| 9 | consts | |
| 13735 | 10 | Factorisation :: ['a::ring, 'a list, 'a] => bool | 
| 7998 | 11 | (* factorisation of x into a list of irred factors and a unit u *) | 
| 12 | ||
| 13 | defs | |
| 14 | Factorisation_def "Factorisation x factors u == | |
| 15 | x = foldr op* factors u & | |
| 13735 | 16 | (ALL a : set factors. irred a) & u dvd 1" | 
| 7998 | 17 | |
| 18 | end |