author | wenzelm |
Mon, 01 Mar 2010 17:12:43 +0100 | |
changeset 35414 | cc8e4276d093 |
parent 29665 | 2b956243d123 |
child 35849 | b5522b51cb1e |
permissions | -rw-r--r-- |
7998 | 1 |
(* |
2 |
Properties of abstract class field |
|
3 |
Author: Clemens Ballarin, started 15 November 1997 |
|
4 |
*) |
|
5 |
||
17479 | 6 |
theory Field imports Factor PID begin |
7998 | 7 |
|
17479 | 8 |
instance field < "domain" |
9 |
apply intro_classes |
|
10 |
apply (rule field_one_not_zero) |
|
11 |
apply (erule field_integral) |
|
12 |
done |
|
7998 | 13 |
|
17479 | 14 |
instance field < factorial |
15 |
apply intro_classes |
|
16 |
apply (erule field_fact_prime) |
|
17 |
done |
|
7998 | 18 |
|
19 |
end |