src/HOL/ex/Lagrange.thy
author krauss
Tue, 28 Sep 2010 09:54:07 +0200
changeset 39754 150f831ce4a3
parent 37885 c43805c80eb6
child 58776 95e58e04e534
permissions -rw-r--r--
no longer declare .psimps rules as [simp]. This regularly caused confusion (e.g., they show up in simp traces when the regular simp rules are disabled). In the few places where the rules are used, explicitly mentioning them actually clarifies the proof text.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11375
paulson
parents: 5078
diff changeset
     1
(*  Title:      HOL/ex/Lagrange.thy
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     2
    Author:     Tobias Nipkow
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     3
    Copyright   1996 TU Muenchen
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     4
*)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     5
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
     6
header {* A lemma for Lagrange's theorem *}
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15069
diff changeset
     8
theory Lagrange imports Main begin
14603
985eb6708207 Moved ring stuff from ex into Ring_and_Field.
nipkow
parents: 11375
diff changeset
     9
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    10
text {* This theory only contains a single theorem, which is a lemma
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    11
in Lagrange's proof that every natural number is the sum of 4 squares.
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    12
Its sole purpose is to demonstrate ordered rewriting for commutative
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    13
rings.
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    14
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    15
The enterprising reader might consider proving all of Lagrange's
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    16
theorem.  *}
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    17
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22173
diff changeset
    18
definition sq :: "'a::times => 'a" where "sq x == x*x"
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    19
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    20
text {* The following lemma essentially shows that every natural
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    21
number is the sum of four squares, provided all prime numbers are.
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    22
However, this is an abstract theorem about commutative rings.  It has,
495c799df31d tuned headers etc.;
wenzelm
parents: 16740
diff changeset
    23
a priori, nothing to do with nat. *}
14603
985eb6708207 Moved ring stuff from ex into Ring_and_Field.
nipkow
parents: 11375
diff changeset
    24
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22173
diff changeset
    25
lemma Lagrange_lemma: fixes x1 :: "'a::comm_ring" shows
20807
wenzelm
parents: 19736
diff changeset
    26
  "(sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) =
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22173
diff changeset
    27
   sq (x1*y1 - x2*y2 - x3*y3 - x4*y4)  +
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22173
diff changeset
    28
   sq (x1*y2 + x2*y1 + x3*y4 - x4*y3)  +
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22173
diff changeset
    29
   sq (x1*y3 - x2*y4 + x3*y1 + x4*y2)  +
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22173
diff changeset
    30
   sq (x1*y4 + x2*y3 - x3*y2 + x4*y1)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 30601
diff changeset
    31
by (simp only: sq_def field_simps)
14603
985eb6708207 Moved ring stuff from ex into Ring_and_Field.
nipkow
parents: 11375
diff changeset
    32
985eb6708207 Moved ring stuff from ex into Ring_and_Field.
nipkow
parents: 11375
diff changeset
    33
25475
d5a382ccb5cc challenge by John Harrison: down to 12s (was 17s, was 75s);
wenzelm
parents: 23477
diff changeset
    34
text {* A challenge by John Harrison. Takes about 12s on a 1.6GHz machine. *}
14603
985eb6708207 Moved ring stuff from ex into Ring_and_Field.
nipkow
parents: 11375
diff changeset
    35
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22173
diff changeset
    36
lemma fixes p1 :: "'a::comm_ring" shows
20807
wenzelm
parents: 19736
diff changeset
    37
  "(sq p1 + sq q1 + sq r1 + sq s1 + sq t1 + sq u1 + sq v1 + sq w1) * 
wenzelm
parents: 19736
diff changeset
    38
   (sq p2 + sq q2 + sq r2 + sq s2 + sq t2 + sq u2 + sq v2 + sq w2) 
wenzelm
parents: 19736
diff changeset
    39
    = sq (p1*p2 - q1*q2 - r1*r2 - s1*s2 - t1*t2 - u1*u2 - v1*v2 - w1*w2) + 
wenzelm
parents: 19736
diff changeset
    40
      sq (p1*q2 + q1*p2 + r1*s2 - s1*r2 + t1*u2 - u1*t2 - v1*w2 + w1*v2) +
wenzelm
parents: 19736
diff changeset
    41
      sq (p1*r2 - q1*s2 + r1*p2 + s1*q2 + t1*v2 + u1*w2 - v1*t2 - w1*u2) +
wenzelm
parents: 19736
diff changeset
    42
      sq (p1*s2 + q1*r2 - r1*q2 + s1*p2 + t1*w2 - u1*v2 + v1*u2 - w1*t2) +
wenzelm
parents: 19736
diff changeset
    43
      sq (p1*t2 - q1*u2 - r1*v2 - s1*w2 + t1*p2 + u1*q2 + v1*r2 + w1*s2) +
wenzelm
parents: 19736
diff changeset
    44
      sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +
wenzelm
parents: 19736
diff changeset
    45
      sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +
wenzelm
parents: 19736
diff changeset
    46
      sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 30601
diff changeset
    47
by (simp only: sq_def field_simps)
14603
985eb6708207 Moved ring stuff from ex into Ring_and_Field.
nipkow
parents: 11375
diff changeset
    48
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    49
end