Theory Groebner_Examples

Up to index of Isabelle/HOL/ex

theory Groebner_Examples
imports Groebner_Basis
begin

(*  Title:      HOL/ex/Groebner_Examples.thy
    ID:         $Id: Groebner_Examples.thy,v 1.6 2008/03/18 19:33:33 wenzelm Exp $
    Author:     Amine Chaieb, TU Muenchen
*)

header {* Groebner Basis Examples *}

theory Groebner_Examples
imports Groebner_Basis
begin

subsection {* Basic examples *}

lemma "3 ^ 3 == (?X::'a::{number_ring,recpower})"
  by sring_norm

lemma "(x - (-2))^5 == ?X::int"
  by sring_norm

lemma "(x - (-2))^5  * (y - 78) ^ 8 == ?X::int"
  by sring_norm

lemma "((-3) ^ (Suc (Suc (Suc 0)))) == (X::'a::{number_ring,recpower})"
  apply (simp only: power_Suc power_0)
  apply (simp only: comp_arith)
  oops

lemma "((x::int) + y)^3 - 1 = (x - z)^2 - 10 ==> x = z + 3 ==> x = - y"
  by algebra

lemma "(4::nat) + 4 = 3 + 5"
  by algebra

lemma "(4::int) + 0 = 4"
  apply algebra?
  by simp

lemma
  assumes "a * x^2 + b * x + c = (0::int)" and "d * x^2 + e * x + f = 0"
  shows "d^2*c^2 - 2*d*c*a*f + a^2*f^2 - e*d*b*c - e*b*a*f + a*e^2*c + f*d*b^2 = 0"
  using assms by algebra

lemma "(x::int)^3  - x^2  - 5*x - 3 = 0 <-> (x = 3 ∨ x = -1)"
  by algebra

theorem "x* (x² - x  - 5) - 3 = (0::int) <-> (x = 3 ∨ x = -1)"
  by algebra

lemma
  fixes x::"'a::{idom,recpower,number_ring}"
  shows "x^2*y = x^2 & x*y^2 = y^2 <->  x=1 & y=1 | x=0 & y=0"
  by algebra

subsection {* Lemmas for Lagrange's theorem *}

definition
  sq :: "'a::times => 'a" where
  "sq x == x*x"

lemma
  fixes x1 :: "'a::{idom,recpower,number_ring}"
  shows
  "(sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) =
    sq (x1*y1 - x2*y2 - x3*y3 - x4*y4)  +
    sq (x1*y2 + x2*y1 + x3*y4 - x4*y3)  +
    sq (x1*y3 - x2*y4 + x3*y1 + x4*y2)  +
    sq (x1*y4 + x2*y3 - x3*y2 + x4*y1)"
  by (algebra add: sq_def)

lemma
  fixes p1 :: "'a::{idom,recpower,number_ring}"
  shows
  "(sq p1 + sq q1 + sq r1 + sq s1 + sq t1 + sq u1 + sq v1 + sq w1) *
   (sq p2 + sq q2 + sq r2 + sq s2 + sq t2 + sq u2 + sq v2 + sq w2)
    = sq (p1*p2 - q1*q2 - r1*r2 - s1*s2 - t1*t2 - u1*u2 - v1*v2 - w1*w2) +
      sq (p1*q2 + q1*p2 + r1*s2 - s1*r2 + t1*u2 - u1*t2 - v1*w2 + w1*v2) +
      sq (p1*r2 - q1*s2 + r1*p2 + s1*q2 + t1*v2 + u1*w2 - v1*t2 - w1*u2) +
      sq (p1*s2 + q1*r2 - r1*q2 + s1*p2 + t1*w2 - u1*v2 + v1*u2 - w1*t2) +
      sq (p1*t2 - q1*u2 - r1*v2 - s1*w2 + t1*p2 + u1*q2 + v1*r2 + w1*s2) +
      sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +
      sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +
      sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)"
  by (algebra add: sq_def)


subsection {* Colinearity is invariant by rotation *}

types point = "int × int"

definition collinear ::"point => point => point => bool" where
  "collinear ≡ λ(Ax,Ay) (Bx,By) (Cx,Cy).
    ((Ax - Bx) * (By - Cy) = (Ay - By) * (Bx - Cx))"

lemma collinear_inv_rotation:
  assumes "collinear (Ax, Ay) (Bx, By) (Cx, Cy)" and "c² + s² = 1"
  shows "collinear (Ax * c - Ay * s, Ay * c + Ax * s)
    (Bx * c - By * s, By * c + Bx * s) (Cx * c - Cy * s, Cy * c + Cx * s)"
  using assms 
  by (algebra add: collinear_def split_def fst_conv snd_conv)

lemma "EX (d::int). a*y - a*x = n*d ==> EX u v. a*u + n*v = 1 ==> EX e. y - x = n*e"
  by algebra

end

Basic examples

lemma

  (3::'a) ^ 3 == 27::'a

lemma

  (x - -2) ^ 5 ==
  x ^ 5 + (10 * x ^ 4 + (40 * x ^ 3 + (80 * x ^ 2 + (80 * x + 32))))

lemma

  (x - -2) ^ 5 * (y - 78) ^ 8 ==
  x ^ 5 * y ^ 8 +
  (-624 * (x ^ 5 * y ^ 7) +
   (10 * (x ^ 4 * y ^ 8) +
    (170352 * (x ^ 5 * y ^ 6) +
     (-6240 * (x ^ 4 * y ^ 7) +
      (40 * (x ^ 3 * y ^ 8) +
       (-26574912 * (x ^ 5 * y ^ 5) +
        (1703520 * (x ^ 4 * y ^ 6) +
         (-24960 * (x ^ 3 * y ^ 7) +
          (80 * (x ^ 2 * y ^ 8) +
           (2591053920 * (x ^ 5 * y ^ 4) +
            (-265749120 * (x ^ 4 * y ^ 5) +
             (6814080 * (x ^ 3 * y ^ 6) +
              (-49920 * (x ^ 2 * y ^ 7) +
               (80 * (x * y ^ 8) +
                (-161681764608 * (x ^ 5 * y ^ 3) +
                 (25910539200 * (x ^ 4 * y ^ 4) +
                  (-1062996480 * (x ^ 3 * y ^ 5) +
                   (13628160 * (x ^ 2 * y ^ 6) +
                    (-49920 * (x * y ^ 7) +
                     (32 * y ^ 8 +
                      (6305588819712 * (x ^ 5 * y ^ 2) +
                       (-1616817646080 * (x ^ 4 * y ^ 3) +
                        (103642156800 * (x ^ 3 * y ^ 4) +
                         (-2125992960 * (x ^ 2 * y ^ 5) +
                          (13628160 * (x * y ^ 6) +
                           (-19968 * y ^ 7 +
                            (-140524550839296 * (x ^ 5 * y) +
                             (63055888197120 * (x ^ 4 * y ^ 2) +
                              (-6467270584320 * (x ^ 3 * y ^ 3) +
                               (207284313600 * (x ^ 2 * y ^ 4) +
                                (-2125992960 * (x * y ^ 5) +
                                 (5451264 * y ^ 6 +
                                  (1370114370683136 * x ^ 5 +
                                   (-1405245508392960 * (x ^ 4 * y) +
                                    (252223552788480 * (x ^ 3 * y ^ 2) +
                                     (-12934541168640 * (x ^ 2 * y ^ 3) +
                                      (207284313600 * (x * y ^ 4) +
                                       (-850397184 * y ^ 5 +
                                        (13701143706831360 * x ^ 4 +
                                         (-5620982033571840 * (x ^ 3 * y) +
  (504447105576960 * (x ^ 2 * y ^ 2) +
   (-12934541168640 * (x * y ^ 3) +
    (82913725440 * y ^ 4 +
     (54804574827325440 * x ^ 3 +
      (-11241964067143680 * (x ^ 2 * y) +
       (504447105576960 * (x * y ^ 2) +
        (-5173816467456 * y ^ 3 +
         (109609149654650880 * x ^ 2 +
          (-11241964067143680 * (x * y) +
           (201778842230784 * y ^ 2 +
            (109609149654650880 * x +
             (-4496785626857472 * y +
              43843659861860352))))))))))))))))))))))))))))))))))))))))))))))))))))

lemma

  [| (x + y) ^ 3 - 1 = (x - z) ^ 2 - 10; x = z + 3 |] ==> x = - y

lemma

  4 + 4 = 3 + 5

lemma

  4 + 0 = 4

lemma

  [| a * x ^ 2 + b * x + c = 0; d * x ^ 2 + e * x + f = 0 |]
  ==> d ^ 2 * c ^ 2 - 2 * d * c * a * f + a ^ 2 * f ^ 2 - e * d * b * c -
      e * b * a * f +
      a * e ^ 2 * c +
      f * d * b ^ 2 =
      0

lemma

  (x ^ 3 - x ^ 2 - 5 * x - 3 = 0) = (x = 3 ∨ x = -1)

theorem

  (x * (x ^ 2 - x - 5) - 3 = 0) = (x = 3 ∨ x = -1)

lemma

  (x ^ 2 * y = x ^ 2 ∧ x * y ^ 2 = y ^ 2) =
  (x = (1::'a) ∧ y = (1::'a) ∨ x = (0::'a) ∧ y = (0::'a))

Lemmas for Lagrange's theorem

lemma

  (sq x1.0 + sq x2.0 + sq x3.0 + sq x4.0) *
  (sq y1.0 + sq y2.0 + sq y3.0 + sq y4.0) =
  sq (x1.0 * y1.0 - x2.0 * y2.0 - x3.0 * y3.0 - x4.0 * y4.0) +
  sq (x1.0 * y2.0 + x2.0 * y1.0 + x3.0 * y4.0 - x4.0 * y3.0) +
  sq (x1.0 * y3.0 - x2.0 * y4.0 + x3.0 * y1.0 + x4.0 * y2.0) +
  sq (x1.0 * y4.0 + x2.0 * y3.0 - x3.0 * y2.0 + x4.0 * y1.0)

lemma

  (sq p1.0 + sq q1.0 + sq r1.0 + sq s1.0 + sq t1.0 + sq u1.0 + sq v1.0 +
   sq w1.0) *
  (sq p2.0 + sq q2.0 + sq r2.0 + sq s2.0 + sq t2.0 + sq u2.0 + sq v2.0 +
   sq w2.0) =
  sq (p1.0 * p2.0 - q1.0 * q2.0 - r1.0 * r2.0 - s1.0 * s2.0 - t1.0 * t2.0 -
      u1.0 * u2.0 -
      v1.0 * v2.0 -
      w1.0 * w2.0) +
  sq (p1.0 * q2.0 + q1.0 * p2.0 + r1.0 * s2.0 - s1.0 * r2.0 + t1.0 * u2.0 -
      u1.0 * t2.0 -
      v1.0 * w2.0 +
      w1.0 * v2.0) +
  sq (p1.0 * r2.0 - q1.0 * s2.0 + r1.0 * p2.0 + s1.0 * q2.0 + t1.0 * v2.0 +
      u1.0 * w2.0 -
      v1.0 * t2.0 -
      w1.0 * u2.0) +
  sq (p1.0 * s2.0 + q1.0 * r2.0 - r1.0 * q2.0 + s1.0 * p2.0 + t1.0 * w2.0 -
      u1.0 * v2.0 +
      v1.0 * u2.0 -
      w1.0 * t2.0) +
  sq (p1.0 * t2.0 - q1.0 * u2.0 - r1.0 * v2.0 - s1.0 * w2.0 + t1.0 * p2.0 +
      u1.0 * q2.0 +
      v1.0 * r2.0 +
      w1.0 * s2.0) +
  sq (p1.0 * u2.0 + q1.0 * t2.0 - r1.0 * w2.0 + s1.0 * v2.0 - t1.0 * q2.0 +
      u1.0 * p2.0 -
      v1.0 * s2.0 +
      w1.0 * r2.0) +
  sq (p1.0 * v2.0 + q1.0 * w2.0 + r1.0 * t2.0 - s1.0 * u2.0 - t1.0 * r2.0 +
      u1.0 * s2.0 +
      v1.0 * p2.0 -
      w1.0 * q2.0) +
  sq (p1.0 * w2.0 - q1.0 * v2.0 + r1.0 * u2.0 + s1.0 * t2.0 - t1.0 * s2.0 -
      u1.0 * r2.0 +
      v1.0 * q2.0 +
      w1.0 * p2.0)

Colinearity is invariant by rotation

lemma collinear_inv_rotation:

  [| collinear (Ax, Ay) (Bx, By) (Cx, Cy); c ^ 2 + s ^ 2 = 1 |]
  ==> collinear (Ax * c - Ay * s, Ay * c + Ax * s)
       (Bx * c - By * s, By * c + Bx * s) (Cx * c - Cy * s, Cy * c + Cx * s)

lemma

  [| ∃d. a * y - a * x = n * d; ∃u v. a * u + n * v = 1 |] ==> ∃e. y - x = n * e