src/HOL/Integ/Lagrange.ML
author paulson
Tue, 01 Jul 1997 17:34:42 +0200
changeset 3477 3aced7fa7d8b
parent 3239 6e2ceb50e17b
child 4230 eb5586526bc9
permissions -rw-r--r--
New theorem priK_inj_eq, injectivity of priK
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2281
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Integ/Lagrange.ML
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     4
    Copyright   1996 TU Muenchen
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     5
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     6
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     7
The following lemma essentially shows that all composite natural numbers are
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     8
sums of fours squares, provided all prime numbers are. However, this is an
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
     9
abstract thm about commutative rings and has a priori nothing to do with nat.
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    10
*)
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    11
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    12
goalw Lagrange.thy [Lagrange.sq_def] "!!x1::'a::cring. \
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    13
\  (sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) = \
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    14
\  sq(x1*y1 - x2*y2 - x3*y3 - x4*y4)  + \
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    15
\  sq(x1*y2 + x2*y1 + x3*y4 - x4*y3)  + \
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    16
\  sq(x1*y3 - x2*y4 + x3*y1 + x4*y2)  + \
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    17
\  sq(x1*y4 + x2*y3 - x3*y2 + x4*y1)";
3239
6e2ceb50e17b Added comment: it is SLOW
paulson
parents: 2281
diff changeset
    18
(*Takes up to three minutes...*)
6e2ceb50e17b Added comment: it is SLOW
paulson
parents: 2281
diff changeset
    19
by (cring_simp 1);
2281
e00c13a29eda Ring Theory.
nipkow
parents:
diff changeset
    20
qed "Lagrange_lemma";