author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 17479 | 68a7acb5f22e |
child 20318 | 0e0ea63fe768 |
permissions | -rw-r--r-- |
7998 | 1 |
(* |
2 |
Factorisation within a factorial domain |
|
3 |
$Id$ |
|
4 |
Author: Clemens Ballarin, started 25 November 1997 |
|
5 |
*) |
|
6 |
||
17479 | 7 |
theory Factor imports Ring begin |
7998 | 8 |
|
17479 | 9 |
constdefs |
10 |
Factorisation :: "['a::ring, 'a list, 'a] => bool" |
|
7998 | 11 |
(* factorisation of x into a list of irred factors and a unit u *) |
17479 | 12 |
"Factorisation x factors u == |
13 |
x = foldr op* factors u & |
|
14 |
(ALL a : set factors. irred a) & u dvd 1" |
|
7998 | 15 |
|
16 |
end |