src/Tools/isac/Knowledge/PolyEq.thy
author Walther Neuper <neuper@ist.tugraz.at>
Tue, 31 Aug 2010 16:38:22 +0200
branchisac-update-Isa09-2
changeset 37967 bd4f7a35e892
parent 37966 78938fc8e022
child 37969 81922154e742
permissions -rw-r--r--
updating Knowledge/Simplify, changes ahead + in test

# overwritelnthy thy --> overwritelnthy @{theory}
# test: thms-replace-Isa02-Isa09-2.sml @{thm ..}
     1 (* theory collecting all knowledge 
     2    (predicates 'is_rootEq_in', 'is_sqrt_in', 'is_ratEq_in')
     3    for PolynomialEquations.
     4    alternative dependencies see Isac.thy
     5    created by: rlang 
     6          date: 02.07
     7    changed by: rlang
     8    last change by: rlang
     9              date: 03.06.03
    10    (c) by Richard Lang, 2003
    11 *)
    12 
    13 theory PolyEq imports LinEq RootRatEq begin 
    14 
    15 consts
    16 
    17 (*---------scripts--------------------------*)
    18   Complete'_square
    19              :: "[bool,real, 
    20 		   bool list] => bool list"
    21                ("((Script Complete'_square (_ _ =))// 
    22                   (_))" 9)
    23  (*----- poly ----- *)	 
    24   Normalize'_poly
    25              :: "[bool,real, 
    26 		   bool list] => bool list"
    27                ("((Script Normalize'_poly (_ _=))// 
    28                   (_))" 9)
    29   Solve'_d0'_polyeq'_equation
    30              :: "[bool,real, 
    31 		   bool list] => bool list"
    32                ("((Script Solve'_d0'_polyeq'_equation (_ _ =))// 
    33                   (_))" 9)
    34   Solve'_d1'_polyeq'_equation
    35              :: "[bool,real, 
    36 		   bool list] => bool list"
    37                ("((Script Solve'_d1'_polyeq'_equation (_ _ =))// 
    38                   (_))" 9)
    39   Solve'_d2'_polyeq'_equation
    40              :: "[bool,real, 
    41 		   bool list] => bool list"
    42                ("((Script Solve'_d2'_polyeq'_equation (_ _ =))// 
    43                   (_))" 9)
    44   Solve'_d2'_polyeq'_sqonly'_equation
    45              :: "[bool,real, 
    46 		   bool list] => bool list"
    47                ("((Script Solve'_d2'_polyeq'_sqonly'_equation (_ _ =))// 
    48                   (_))" 9)
    49   Solve'_d2'_polyeq'_bdvonly'_equation
    50              :: "[bool,real, 
    51 		   bool list] => bool list"
    52                ("((Script Solve'_d2'_polyeq'_bdvonly'_equation (_ _ =))// 
    53                   (_))" 9)
    54   Solve'_d2'_polyeq'_pq'_equation
    55              :: "[bool,real, 
    56 		   bool list] => bool list"
    57                ("((Script Solve'_d2'_polyeq'_pq'_equation (_ _ =))// 
    58                   (_))" 9)
    59   Solve'_d2'_polyeq'_abc'_equation
    60              :: "[bool,real, 
    61 		   bool list] => bool list"
    62                ("((Script Solve'_d2'_polyeq'_abc'_equation (_ _ =))// 
    63                   (_))" 9)
    64   Solve'_d3'_polyeq'_equation
    65              :: "[bool,real, 
    66 		   bool list] => bool list"
    67                ("((Script Solve'_d3'_polyeq'_equation (_ _ =))// 
    68                   (_))" 9)
    69   Solve'_d4'_polyeq'_equation
    70              :: "[bool,real, 
    71 		   bool list] => bool list"
    72                ("((Script Solve'_d4'_polyeq'_equation (_ _ =))// 
    73                   (_))" 9)
    74   Biquadrat'_poly
    75              :: "[bool,real, 
    76 		   bool list] => bool list"
    77                ("((Script Biquadrat'_poly (_ _=))// 
    78                   (_))" 9)
    79 
    80 (*-------------------- rules -------------------------------------------------*)
    81 axioms 
    82 
    83   cancel_leading_coeff1 "Not (c =!= 0) ==> (a + b*bdv + c*bdv^^^2 = 0) = 
    84 			                   (a/c + b/c*bdv + bdv^^^2 = 0)"
    85   cancel_leading_coeff2 "Not (c =!= 0) ==> (a - b*bdv + c*bdv^^^2 = 0) = 
    86 			                   (a/c - b/c*bdv + bdv^^^2 = 0)"
    87   cancel_leading_coeff3 "Not (c =!= 0) ==> (a + b*bdv - c*bdv^^^2 = 0) = 
    88 			                   (a/c + b/c*bdv - bdv^^^2 = 0)"
    89 
    90   cancel_leading_coeff4 "Not (c =!= 0) ==> (a +   bdv + c*bdv^^^2 = 0) = 
    91 			                   (a/c + 1/c*bdv + bdv^^^2 = 0)"
    92   cancel_leading_coeff5 "Not (c =!= 0) ==> (a -   bdv + c*bdv^^^2 = 0) = 
    93 			                   (a/c - 1/c*bdv + bdv^^^2 = 0)"
    94   cancel_leading_coeff6 "Not (c =!= 0) ==> (a +   bdv - c*bdv^^^2 = 0) = 
    95 			                   (a/c + 1/c*bdv - bdv^^^2 = 0)"
    96 
    97   cancel_leading_coeff7 "Not (c =!= 0) ==> (    b*bdv + c*bdv^^^2 = 0) = 
    98 			                   (    b/c*bdv + bdv^^^2 = 0)"
    99   cancel_leading_coeff8 "Not (c =!= 0) ==> (    b*bdv - c*bdv^^^2 = 0) = 
   100 			                   (    b/c*bdv - bdv^^^2 = 0)"
   101 
   102   cancel_leading_coeff9 "Not (c =!= 0) ==> (      bdv + c*bdv^^^2 = 0) = 
   103 			                   (      1/c*bdv + bdv^^^2 = 0)"
   104   cancel_leading_coeff10"Not (c =!= 0) ==> (      bdv - c*bdv^^^2 = 0) = 
   105 			                   (      1/c*bdv - bdv^^^2 = 0)"
   106 
   107   cancel_leading_coeff11"Not (c =!= 0) ==> (a +      b*bdv^^^2 = 0) = 
   108 			                   (a/b +      bdv^^^2 = 0)"
   109   cancel_leading_coeff12"Not (c =!= 0) ==> (a -      b*bdv^^^2 = 0) = 
   110 			                   (a/b -      bdv^^^2 = 0)"
   111   cancel_leading_coeff13"Not (c =!= 0) ==> (         b*bdv^^^2 = 0) = 
   112 			                   (           bdv^^^2 = 0/b)"
   113 
   114   complete_square1      "(q + p*bdv + bdv^^^2 = 0) = 
   115 		         (q + (p/2 + bdv)^^^2 = (p/2)^^^2)"
   116   complete_square2      "(    p*bdv + bdv^^^2 = 0) = 
   117 		         (    (p/2 + bdv)^^^2 = (p/2)^^^2)"
   118   complete_square3      "(      bdv + bdv^^^2 = 0) = 
   119 		         (    (1/2 + bdv)^^^2 = (1/2)^^^2)"
   120 		        
   121   complete_square4      "(q - p*bdv + bdv^^^2 = 0) = 
   122 		         (q + (p/2 - bdv)^^^2 = (p/2)^^^2)"
   123   complete_square5      "(q + p*bdv - bdv^^^2 = 0) = 
   124 		         (q + (p/2 - bdv)^^^2 = (p/2)^^^2)"
   125 
   126   square_explicit1      "(a + b^^^2 = c) = ( b^^^2 = c - a)"
   127   square_explicit2      "(a - b^^^2 = c) = (-(b^^^2) = c - a)"
   128 
   129   bdv_explicit1         "(a + bdv = b) = (bdv = - a + b)"
   130   bdv_explicit2         "(a - bdv = b) = ((-1)*bdv = - a + b)"
   131   bdv_explicit3         "((-1)*bdv = b) = (bdv = (-1)*b)"
   132 
   133   plus_leq              "(0 <= a + b) = ((-1)*b <= a)"(*Isa?*)
   134   minus_leq             "(0 <= a - b) = (     b <= a)"(*Isa?*)
   135 
   136 (*-- normalize --*)
   137   (*WN0509 compare LinEq.all_left "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)"*)
   138   all_left              "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"
   139   makex1_x              "a^^^1  = a"  
   140   real_assoc_1          "a+(b+c) = a+b+c"
   141   real_assoc_2          "a*(b*c) = a*b*c"
   142 
   143 (* ---- degree 0 ----*)
   144   d0_true               "(0=0) = True"
   145   d0_false              "[|Not(bdv occurs_in a);Not(a=0)|] ==> (a=0) = False"
   146 (* ---- degree 1 ----*)
   147   d1_isolate_add1
   148    "[|Not(bdv occurs_in a)|] ==> (a + b*bdv = 0) = (b*bdv = (-1)*a)"
   149   d1_isolate_add2
   150    "[|Not(bdv occurs_in a)|] ==> (a +   bdv = 0) = (  bdv = (-1)*a)"
   151   d1_isolate_div
   152    "[|Not(b=0);Not(bdv occurs_in c)|] ==> (b*bdv = c) = (bdv = c/b)"
   153 (* ---- degree 2 ----*)
   154   d2_isolate_add1
   155    "[|Not(bdv occurs_in a)|] ==> (a + b*bdv^^^2=0) = (b*bdv^^^2= (-1)*a)"
   156   d2_isolate_add2
   157    "[|Not(bdv occurs_in a)|] ==> (a +   bdv^^^2=0) = (  bdv^^^2= (-1)*a)"
   158   d2_isolate_div
   159    "[|Not(b=0);Not(bdv occurs_in c)|] ==> (b*bdv^^^2=c) = (bdv^^^2=c/b)"
   160 
   161   d2_prescind1          "(a*bdv + b*bdv^^^2 = 0) = (bdv*(a +b*bdv)=0)"
   162   d2_prescind2          "(a*bdv +   bdv^^^2 = 0) = (bdv*(a +  bdv)=0)"
   163   d2_prescind3          "(  bdv + b*bdv^^^2 = 0) = (bdv*(1+b*bdv)=0)"
   164   d2_prescind4          "(  bdv +   bdv^^^2 = 0) = (bdv*(1+  bdv)=0)"
   165   (* eliminate degree 2 *)
   166   (* thm for neg arguments in sqroot have postfix _neg *)
   167   d2_sqrt_equation1     "[|(0<=c);Not(bdv occurs_in c)|] ==> 
   168                          (bdv^^^2=c) = ((bdv=sqrt c) | (bdv=(-1)*sqrt c ))"
   169   d2_sqrt_equation1_neg
   170   "[|(c<0);Not(bdv occurs_in c)|] ==> (bdv^^^2=c) = False"
   171   d2_sqrt_equation2     "(bdv^^^2=0) = (bdv=0)"
   172   d2_sqrt_equation3     "(b*bdv^^^2=0) = (bdv=0)"
   173   d2_reduce_equation1   "(bdv*(a +b*bdv)=0) = ((bdv=0)|(a+b*bdv=0))"
   174   d2_reduce_equation2   "(bdv*(a +  bdv)=0) = ((bdv=0)|(a+  bdv=0))"
   175   d2_pqformula1         "[|0<=p^^^2 - 4*q|] ==> (q+p*bdv+   bdv^^^2=0) =
   176                            ((bdv= (-1)*(p/2) + sqrt(p^^^2 - 4*q)/2) 
   177                           | (bdv= (-1)*(p/2) - sqrt(p^^^2 - 4*q)/2))"
   178   d2_pqformula1_neg     "[|p^^^2 - 4*q<0|] ==> (q+p*bdv+   bdv^^^2=0) = False"
   179   d2_pqformula2         "[|0<=p^^^2 - 4*q|] ==> (q+p*bdv+1*bdv^^^2=0) = 
   180                            ((bdv= (-1)*(p/2) + sqrt(p^^^2 - 4*q)/2) 
   181                           | (bdv= (-1)*(p/2) - sqrt(p^^^2 - 4*q)/2))"
   182   d2_pqformula2_neg     "[|p^^^2 - 4*q<0|] ==> (q+p*bdv+1*bdv^^^2=0) = False"
   183   d2_pqformula3         "[|0<=1 - 4*q|] ==> (q+  bdv+   bdv^^^2=0) = 
   184                            ((bdv= (-1)*(1/2) + sqrt(1 - 4*q)/2) 
   185                           | (bdv= (-1)*(1/2) - sqrt(1 - 4*q)/2))"
   186   d2_pqformula3_neg     "[|1 - 4*q<0|] ==> (q+  bdv+   bdv^^^2=0) = False"
   187   d2_pqformula4         "[|0<=1 - 4*q|] ==> (q+  bdv+1*bdv^^^2=0) = 
   188                            ((bdv= (-1)*(1/2) + sqrt(1 - 4*q)/2) 
   189                           | (bdv= (-1)*(1/2) - sqrt(1 - 4*q)/2))"
   190   d2_pqformula4_neg     "[|1 - 4*q<0|] ==> (q+  bdv+1*bdv^^^2=0) = False"
   191   d2_pqformula5         "[|0<=p^^^2 - 0|] ==> (  p*bdv+   bdv^^^2=0) =
   192                            ((bdv= (-1)*(p/2) + sqrt(p^^^2 - 0)/2) 
   193                           | (bdv= (-1)*(p/2) - sqrt(p^^^2 - 0)/2))"
   194   (* d2_pqformula5_neg not need p^2 never less zero in R *)
   195   d2_pqformula6         "[|0<=p^^^2 - 0|] ==> (  p*bdv+1*bdv^^^2=0) = 
   196                            ((bdv= (-1)*(p/2) + sqrt(p^^^2 - 0)/2) 
   197                           | (bdv= (-1)*(p/2) - sqrt(p^^^2 - 0)/2))"
   198   (* d2_pqformula6_neg not need p^2 never less zero in R *)
   199   d2_pqformula7         "[|0<=1 - 0|] ==> (    bdv+   bdv^^^2=0) = 
   200                            ((bdv= (-1)*(1/2) + sqrt(1 - 0)/2) 
   201                           | (bdv= (-1)*(1/2) - sqrt(1 - 0)/2))"
   202   (* d2_pqformula7_neg not need, because 1<0 ==> False*)
   203   d2_pqformula8         "[|0<=1 - 0|] ==> (    bdv+1*bdv^^^2=0) = 
   204                            ((bdv= (-1)*(1/2) + sqrt(1 - 0)/2) 
   205                           | (bdv= (-1)*(1/2) - sqrt(1 - 0)/2))"
   206   (* d2_pqformula8_neg not need, because 1<0 ==> False*)
   207   d2_pqformula9         "[|Not(bdv occurs_in q); 0<= (-1)*4*q|] ==> 
   208                            (q+    1*bdv^^^2=0) = ((bdv= 0 + sqrt(0 - 4*q)/2) 
   209                                                 | (bdv= 0 - sqrt(0 - 4*q)/2))"
   210   d2_pqformula9_neg
   211    "[|Not(bdv occurs_in q); (-1)*4*q<0|] ==> (q+    1*bdv^^^2=0) = False"
   212   d2_pqformula10
   213    "[|Not(bdv occurs_in q); 0<= (-1)*4*q|] ==> (q+     bdv^^^2=0) = 
   214            ((bdv= 0 + sqrt(0 - 4*q)/2) 
   215           | (bdv= 0 - sqrt(0 - 4*q)/2))"
   216   d2_pqformula10_neg
   217    "[|Not(bdv occurs_in q); (-1)*4*q<0|] ==> (q+     bdv^^^2=0) = False"
   218   d2_abcformula1
   219    "[|0<=b^^^2 - 4*a*c|] ==> (c + b*bdv+a*bdv^^^2=0) =
   220            ((bdv=( -b + sqrt(b^^^2 - 4*a*c))/(2*a)) 
   221           | (bdv=( -b - sqrt(b^^^2 - 4*a*c))/(2*a)))"
   222   d2_abcformula1_neg
   223    "[|b^^^2 - 4*a*c<0|] ==> (c + b*bdv+a*bdv^^^2=0) = False"
   224   d2_abcformula2
   225    "[|0<=1 - 4*a*c|]     ==> (c+    bdv+a*bdv^^^2=0) = 
   226            ((bdv=( -1 + sqrt(1 - 4*a*c))/(2*a)) 
   227           | (bdv=( -1 - sqrt(1 - 4*a*c))/(2*a)))"
   228   d2_abcformula2_neg
   229    "[|1 - 4*a*c<0|]     ==> (c+    bdv+a*bdv^^^2=0) = False"
   230   d2_abcformula3
   231    "[|0<=b^^^2 - 4*1*c|] ==> (c + b*bdv+  bdv^^^2=0) =
   232            ((bdv=( -b + sqrt(b^^^2 - 4*1*c))/(2*1)) 
   233           | (bdv=( -b - sqrt(b^^^2 - 4*1*c))/(2*1)))"
   234   d2_abcformula3_neg
   235    "[|b^^^2 - 4*1*c<0|] ==> (c + b*bdv+  bdv^^^2=0) = False"
   236   d2_abcformula4
   237    "[|0<=1 - 4*1*c|] ==> (c +   bdv+  bdv^^^2=0) =
   238            ((bdv=( -1 + sqrt(1 - 4*1*c))/(2*1)) 
   239           | (bdv=( -1 - sqrt(1 - 4*1*c))/(2*1)))"
   240   d2_abcformula4_neg
   241    "[|1 - 4*1*c<0|] ==> (c +   bdv+  bdv^^^2=0) = False"
   242   d2_abcformula5
   243    "[|Not(bdv occurs_in c); 0<=0 - 4*a*c|] ==> (c +  a*bdv^^^2=0) =
   244            ((bdv=( 0 + sqrt(0 - 4*a*c))/(2*a)) 
   245           | (bdv=( 0 - sqrt(0 - 4*a*c))/(2*a)))"
   246   d2_abcformula5_neg
   247    "[|Not(bdv occurs_in c); 0 - 4*a*c<0|] ==> (c +  a*bdv^^^2=0) = False"
   248   d2_abcformula6
   249    "[|Not(bdv occurs_in c); 0<=0 - 4*1*c|]     ==> (c+    bdv^^^2=0) = 
   250            ((bdv=( 0 + sqrt(0 - 4*1*c))/(2*1)) 
   251           | (bdv=( 0 - sqrt(0 - 4*1*c))/(2*1)))"
   252   d2_abcformula6_neg
   253    "[|Not(bdv occurs_in c); 0 - 4*1*c<0|]     ==> (c+    bdv^^^2=0) = False"
   254   d2_abcformula7
   255    "[|0<=b^^^2 - 0|]     ==> (    b*bdv+a*bdv^^^2=0) = 
   256            ((bdv=( -b + sqrt(b^^^2 - 0))/(2*a)) 
   257           | (bdv=( -b - sqrt(b^^^2 - 0))/(2*a)))"
   258   (* d2_abcformula7_neg not need b^2 never less zero in R *)
   259   d2_abcformula8
   260    "[|0<=b^^^2 - 0|] ==> (    b*bdv+  bdv^^^2=0) =
   261            ((bdv=( -b + sqrt(b^^^2 - 0))/(2*1)) 
   262           | (bdv=( -b - sqrt(b^^^2 - 0))/(2*1)))"
   263   (* d2_abcformula8_neg not need b^2 never less zero in R *)
   264   d2_abcformula9
   265    "[|0<=1 - 0|]     ==> (      bdv+a*bdv^^^2=0) = 
   266            ((bdv=( -1 + sqrt(1 - 0))/(2*a)) 
   267           | (bdv=( -1 - sqrt(1 - 0))/(2*a)))"
   268   (* d2_abcformula9_neg not need, because 1<0 ==> False*)
   269   d2_abcformula10
   270    "[|0<=1 - 0|] ==> (      bdv+  bdv^^^2=0) =
   271            ((bdv=( -1 + sqrt(1 - 0))/(2*1)) 
   272           | (bdv=( -1 - sqrt(1 - 0))/(2*1)))"
   273   (* d2_abcformula10_neg not need, because 1<0 ==> False*)
   274 
   275 (* ---- degree 3 ----*)
   276   d3_reduce_equation1
   277   "(a*bdv + b*bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (a + b*bdv + c*bdv^^^2=0))"
   278   d3_reduce_equation2
   279   "(  bdv + b*bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (1 + b*bdv + c*bdv^^^2=0))"
   280   d3_reduce_equation3
   281   "(a*bdv +   bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (a +   bdv + c*bdv^^^2=0))"
   282   d3_reduce_equation4
   283   "(  bdv +   bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (1 +   bdv + c*bdv^^^2=0))"
   284   d3_reduce_equation5
   285   "(a*bdv + b*bdv^^^2 +   bdv^^^3=0) = (bdv=0 | (a + b*bdv +   bdv^^^2=0))"
   286   d3_reduce_equation6
   287   "(  bdv + b*bdv^^^2 +   bdv^^^3=0) = (bdv=0 | (1 + b*bdv +   bdv^^^2=0))"
   288   d3_reduce_equation7
   289   "(a*bdv +   bdv^^^2 +   bdv^^^3=0) = (bdv=0 | (1 +   bdv +   bdv^^^2=0))"
   290   d3_reduce_equation8
   291   "(  bdv +   bdv^^^2 +   bdv^^^3=0) = (bdv=0 | (1 +   bdv +   bdv^^^2=0))"
   292   d3_reduce_equation9
   293   "(a*bdv             + c*bdv^^^3=0) = (bdv=0 | (a         + c*bdv^^^2=0))"
   294   d3_reduce_equation10
   295   "(  bdv             + c*bdv^^^3=0) = (bdv=0 | (1         + c*bdv^^^2=0))"
   296   d3_reduce_equation11
   297   "(a*bdv             +   bdv^^^3=0) = (bdv=0 | (a         +   bdv^^^2=0))"
   298   d3_reduce_equation12
   299   "(  bdv             +   bdv^^^3=0) = (bdv=0 | (1         +   bdv^^^2=0))"
   300   d3_reduce_equation13
   301   "(        b*bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (    b*bdv + c*bdv^^^2=0))"
   302   d3_reduce_equation14
   303   "(          bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (      bdv + c*bdv^^^2=0))"
   304   d3_reduce_equation15
   305   "(        b*bdv^^^2 +   bdv^^^3=0) = (bdv=0 | (    b*bdv +   bdv^^^2=0))"
   306   d3_reduce_equation16
   307   "(          bdv^^^2 +   bdv^^^3=0) = (bdv=0 | (      bdv +   bdv^^^2=0))"
   308   d3_isolate_add1
   309   "[|Not(bdv occurs_in a)|] ==> (a + b*bdv^^^3=0) = (b*bdv^^^3= (-1)*a)"
   310   d3_isolate_add2
   311   "[|Not(bdv occurs_in a)|] ==> (a +   bdv^^^3=0) = (  bdv^^^3= (-1)*a)"
   312   d3_isolate_div
   313    "[|Not(b=0);Not(bdv occurs_in a)|] ==> (b*bdv^^^3=c) = (bdv^^^3=c/b)"
   314   d3_root_equation2
   315   "(bdv^^^3=0) = (bdv=0)"
   316   d3_root_equation1
   317   "(bdv^^^3=c) = (bdv = nroot 3 c)"
   318 
   319 (* ---- degree 4 ----*)
   320  (* RL03.FIXME es wir nicht getestet ob u>0 *)
   321  d4_sub_u1
   322  "(c+b*bdv^^^2+a*bdv^^^4=0) =
   323    ((a*u^^^2+b*u+c=0) & (bdv^^^2=u))"
   324 
   325 (* ---- 7.3.02 von Termorder ---- *)
   326 
   327   bdv_collect_1       "l * bdv + m * bdv = (l + m) * bdv"
   328   bdv_collect_2       "bdv + m * bdv = (1 + m) * bdv"
   329   bdv_collect_3       "l * bdv + bdv = (l + 1) * bdv"
   330 
   331 (*  bdv_collect_assoc0_1 "l * bdv + m * bdv + k = (l + m) * bdv + k"
   332     bdv_collect_assoc0_2 "bdv + m * bdv + k = (1 + m) * bdv + k"
   333     bdv_collect_assoc0_3 "l * bdv + bdv + k = (l + 1) * bdv + k"
   334 *)
   335   bdv_collect_assoc1_1 "l * bdv + (m * bdv + k) = (l + m) * bdv + k"
   336   bdv_collect_assoc1_2 "bdv + (m * bdv + k) = (1 + m) * bdv + k"
   337   bdv_collect_assoc1_3 "l * bdv + (bdv + k) = (l + 1) * bdv + k"
   338 
   339   bdv_collect_assoc2_1 "k + l * bdv + m * bdv = k + (l + m) * bdv"
   340   bdv_collect_assoc2_2 "k + bdv + m * bdv = k + (1 + m) * bdv"
   341   bdv_collect_assoc2_3 "k + l * bdv + bdv = k + (l + 1) * bdv"
   342 
   343 
   344   bdv_n_collect_1      "l * bdv^^^n + m * bdv^^^n = (l + m) * bdv^^^n"
   345   bdv_n_collect_2      " bdv^^^n + m * bdv^^^n = (1 + m) * bdv^^^n"
   346   bdv_n_collect_3      "l * bdv^^^n + bdv^^^n = (l + 1) * bdv^^^n"   (*order!*)
   347 
   348   bdv_n_collect_assoc1_1 "l * bdv^^^n + (m * bdv^^^n + k) = (l + m) * bdv^^^n + k"
   349   bdv_n_collect_assoc1_2 "bdv^^^n + (m * bdv^^^n + k) = (1 + m) * bdv^^^n + k"
   350   bdv_n_collect_assoc1_3 "l * bdv^^^n + (bdv^^^n + k) = (l + 1) * bdv^^^n + k"
   351 
   352   bdv_n_collect_assoc2_1 "k + l * bdv^^^n + m * bdv^^^n = k + (l + m) * bdv^^^n"
   353   bdv_n_collect_assoc2_2 "k + bdv^^^n + m * bdv^^^n = k + (1 + m) * bdv^^^n"
   354   bdv_n_collect_assoc2_3 "k + l * bdv^^^n + bdv^^^n = k + (l + 1) * bdv^^^n"
   355 
   356 (*WN.14.3.03*)
   357   real_minus_div         "- (a / b) = (-1 * a) / b"
   358 
   359   separate_bdv           "(a * bdv) / b = (a / b) * bdv"
   360   separate_bdv_n         "(a * bdv ^^^ n) / b = (a / b) * bdv ^^^ n"
   361   separate_1_bdv         "bdv / b = (1 / b) * bdv"
   362   separate_1_bdv_n       "bdv ^^^ n / b = (1 / b) * bdv ^^^ n"
   363 
   364 ML {*
   365 (*-------------------------rulse-------------------------*)
   366 val PolyEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
   367   append_rls "PolyEq_prls" e_rls 
   368 	     [Calc ("Atools.ident",eval_ident "#ident_"),
   369 	      Calc ("Tools.matches",eval_matches ""),
   370 	      Calc ("Tools.lhs"    ,eval_lhs ""),
   371 	      Calc ("Tools.rhs"    ,eval_rhs ""),
   372 	      Calc ("Poly.is'_expanded'_in",eval_is_expanded_in ""),
   373 	      Calc ("Poly.is'_poly'_in",eval_is_poly_in ""),
   374 	      Calc ("Poly.has'_degree'_in",eval_has_degree_in ""),    
   375               Calc ("Poly.is'_polyrat'_in",eval_is_polyrat_in ""),
   376 	      (*Calc ("Atools.occurs'_in",eval_occurs_in ""),   *) 
   377 	      (*Calc ("Atools.is'_const",eval_const "#is_const_"),*)
   378 	      Calc ("op =",eval_equal "#equal_"),
   379               Calc ("RootEq.is'_rootTerm'_in",eval_is_rootTerm_in ""),
   380 	      Calc ("RatEq.is'_ratequation'_in",eval_is_ratequation_in ""),
   381 	      Thm ("not_true",num_str @{not_true),
   382 	      Thm ("not_false",num_str @{not_false),
   383 	      Thm ("and_true",num_str @{and_true),
   384 	      Thm ("and_false",num_str @{and_false),
   385 	      Thm ("or_true",num_str @{or_true),
   386 	      Thm ("or_false",num_str @{or_false)
   387 	       ];
   388 
   389 val PolyEq_erls = 
   390     merge_rls "PolyEq_erls" LinEq_erls
   391     (append_rls "ops_preds" calculate_Rational
   392 		[Calc ("op =",eval_equal "#equal_"),
   393 		 Thm ("plus_leq", num_str @{plus_leq),
   394 		 Thm ("minus_leq", num_str @{minus_leq),
   395 		 Thm ("rat_leq1", num_str @{rat_leq1),
   396 		 Thm ("rat_leq2", num_str @{rat_leq2),
   397 		 Thm ("rat_leq3", num_str @{rat_leq3)
   398 		 ]);
   399 
   400 val PolyEq_crls = 
   401     merge_rls "PolyEq_crls" LinEq_crls
   402     (append_rls "ops_preds" calculate_Rational
   403 		[Calc ("op =",eval_equal "#equal_"),
   404 		 Thm ("plus_leq", num_str @{plus_leq),
   405 		 Thm ("minus_leq", num_str @{minus_leq),
   406 		 Thm ("rat_leq1", num_str @{rat_leq1),
   407 		 Thm ("rat_leq2", num_str @{rat_leq2),
   408 		 Thm ("rat_leq3", num_str @{rat_leq3)
   409 		 ]);
   410 
   411 val cancel_leading_coeff = prep_rls(
   412   Rls {id = "cancel_leading_coeff", preconds = [], 
   413        rew_ord = ("e_rew_ord",e_rew_ord),
   414       erls = PolyEq_erls, srls = Erls, calc = [], (*asm_thm = [],*)
   415       rules = [Thm ("cancel_leading_coeff1",num_str @{cancel_leading_coeff1),
   416 	       Thm ("cancel_leading_coeff2",num_str @{cancel_leading_coeff2),
   417 	       Thm ("cancel_leading_coeff3",num_str @{cancel_leading_coeff3),
   418 	       Thm ("cancel_leading_coeff4",num_str @{cancel_leading_coeff4),
   419 	       Thm ("cancel_leading_coeff5",num_str @{cancel_leading_coeff5),
   420 	       Thm ("cancel_leading_coeff6",num_str @{cancel_leading_coeff6),
   421 	       Thm ("cancel_leading_coeff7",num_str @{cancel_leading_coeff7),
   422 	       Thm ("cancel_leading_coeff8",num_str @{cancel_leading_coeff8),
   423 	       Thm ("cancel_leading_coeff9",num_str @{cancel_leading_coeff9),
   424 	       Thm ("cancel_leading_coeff10",num_str @{cancel_leading_coeff10),
   425 	       Thm ("cancel_leading_coeff11",num_str @{cancel_leading_coeff11),
   426 	       Thm ("cancel_leading_coeff12",num_str @{cancel_leading_coeff12),
   427 	       Thm ("cancel_leading_coeff13",num_str @{cancel_leading_coeff13)
   428 	       ],
   429       scr = Script ((term_of o the o (parse thy)) 
   430       "empty_script")
   431       }:rls);
   432 
   433 val complete_square = prep_rls(
   434   Rls {id = "complete_square", preconds = [], 
   435        rew_ord = ("e_rew_ord",e_rew_ord),
   436       erls = PolyEq_erls, srls = Erls, calc = [], (*asm_thm = [],*)
   437       rules = [Thm ("complete_square1",num_str @{complete_square1),
   438 	       Thm ("complete_square2",num_str @{complete_square2),
   439 	       Thm ("complete_square3",num_str @{complete_square3),
   440 	       Thm ("complete_square4",num_str @{complete_square4),
   441 	       Thm ("complete_square5",num_str @{complete_square5)
   442 	       ],
   443       scr = Script ((term_of o the o (parse thy)) 
   444       "empty_script")
   445       }:rls);
   446 
   447 val polyeq_simplify = prep_rls(
   448   Rls {id = "polyeq_simplify", preconds = [], 
   449        rew_ord = ("termlessI",termlessI), 
   450        erls = PolyEq_erls, 
   451        srls = Erls, 
   452        calc = [], 
   453        (*asm_thm = [],*)
   454        rules = [Thm  ("real_assoc_1",num_str @{real_assoc_1),
   455 		Thm  ("real_assoc_2",num_str @{real_assoc_2),
   456 		Thm  ("real_diff_minus",num_str @{real_diff_minus),
   457 		Thm  ("real_unari_minus",num_str @{real_unari_minus),
   458 		Thm  ("realpow_multI",num_str @{realpow_multI),
   459 		Calc ("op +",eval_binop "#add_"),
   460 		Calc ("op -",eval_binop "#sub_"),
   461 		Calc ("op *",eval_binop "#mult_"),
   462 		Calc ("HOL.divide", eval_cancel "#divide_"),
   463 		Calc ("Root.sqrt",eval_sqrt "#sqrt_"),
   464 		Calc ("Atools.pow" ,eval_binop "#power_"),
   465                 Rls_ reduce_012
   466                 ],
   467        scr = Script ((term_of o the o (parse thy)) "empty_script")
   468        }:rls);
   469 
   470 ruleset' := overwritelthy @{theory} (!ruleset',
   471 		[("cancel_leading_coeff",cancel_leading_coeff),
   472 		 ("complete_square",complete_square),
   473 		 ("PolyEq_erls",PolyEq_erls),(*FIXXXME:del with rls.rls'*)
   474 		 ("polyeq_simplify",polyeq_simplify)]);
   475 
   476 
   477 (* ------------- polySolve ------------------ *)
   478 (* -- d0 -- *)
   479 (*isolate the bound variable in an d0 equation; 'bdv' is a meta-constant*)
   480 val d0_polyeq_simplify = prep_rls(
   481   Rls {id = "d0_polyeq_simplify", preconds = [],
   482        rew_ord = ("e_rew_ord",e_rew_ord),
   483        erls = PolyEq_erls,
   484        srls = Erls, 
   485        calc = [], 
   486        (*asm_thm = [],*)
   487        rules = [Thm("d0_true",num_str @{d0_true),
   488 		Thm("d0_false",num_str @{d0_false)
   489 		],
   490        scr = Script ((term_of o the o (parse thy)) "empty_script")
   491        }:rls);
   492 
   493 (* -- d1 -- *)
   494 (*isolate the bound variable in an d1 equation; 'bdv' is a meta-constant*)
   495 val d1_polyeq_simplify = prep_rls(
   496   Rls {id = "d1_polyeq_simplify", preconds = [],
   497        rew_ord = ("e_rew_ord",e_rew_ord),
   498        erls = PolyEq_erls,
   499        srls = Erls, 
   500        calc = [], 
   501        (*asm_thm = [("d1_isolate_div","")],*)
   502        rules = [
   503 		Thm("d1_isolate_add1",num_str @{d1_isolate_add1), 
   504 		(* a+bx=0 -> bx=-a *)
   505 		Thm("d1_isolate_add2",num_str @{d1_isolate_add2), 
   506 		(* a+ x=0 ->  x=-a *)
   507 		Thm("d1_isolate_div",num_str @{d1_isolate_div)    
   508 		(*   bx=c -> x=c/b *)  
   509 		],
   510        scr = Script ((term_of o the o (parse thy)) "empty_script")
   511        }:rls);
   512 
   513 (* -- d2 -- *)
   514 (* isolate the bound variable in an d2 equation with bdv only; 
   515    'bdv' is a meta-constant*)
   516 val d2_polyeq_bdv_only_simplify = prep_rls(
   517   Rls {id = "d2_polyeq_bdv_only_simplify", preconds = [],
   518        rew_ord = ("e_rew_ord",e_rew_ord),
   519        erls = PolyEq_erls,
   520        srls = Erls, 
   521        calc = [], 
   522        (*asm_thm = [("d2_sqrt_equation1",""),("d2_sqrt_equation1_neg",""),
   523                   ("d2_isolate_div","")],*)
   524        rules = [Thm("d2_prescind1",num_str @{d2_prescind1),
   525                 (*   ax+bx^2=0 -> x(a+bx)=0 *)
   526 		Thm("d2_prescind2",num_str @{d2_prescind2),
   527                 (*   ax+ x^2=0 -> x(a+ x)=0 *)
   528 		Thm("d2_prescind3",num_str @{d2_prescind3),
   529                 (*    x+bx^2=0 -> x(1+bx)=0 *)
   530 		Thm("d2_prescind4",num_str @{d2_prescind4),
   531                 (*    x+ x^2=0 -> x(1+ x)=0 *)
   532 		Thm("d2_sqrt_equation1",num_str @{d2_sqrt_equation1),
   533                 (* x^2=c   -> x=+-sqrt(c)*)
   534 		Thm("d2_sqrt_equation1_neg",num_str @{d2_sqrt_equation1_neg),
   535                 (* [0<c] x^2=c  -> [] *)
   536 		Thm("d2_sqrt_equation2",num_str @{d2_sqrt_equation2),
   537                 (*  x^2=0 ->    x=0    *)
   538 		Thm("d2_reduce_equation1",num_str @{d2_reduce_equation1),
   539                 (* x(a+bx)=0 -> x=0 | a+bx=0*)
   540 		Thm("d2_reduce_equation2",num_str @{d2_reduce_equation2),
   541                 (* x(a+ x)=0 -> x=0 | a+ x=0*)
   542 		Thm("d2_isolate_div",num_str @{d2_isolate_div)
   543                 (* bx^2=c -> x^2=c/b*)
   544 		],
   545        scr = Script ((term_of o the o (parse thy)) "empty_script")
   546        }:rls);
   547 
   548 (* isolate the bound variable in an d2 equation with sqrt only; 
   549    'bdv' is a meta-constant*)
   550 val d2_polyeq_sq_only_simplify = prep_rls(
   551   Rls {id = "d2_polyeq_sq_only_simplify", preconds = [],
   552        rew_ord = ("e_rew_ord",e_rew_ord),
   553        erls = PolyEq_erls,
   554        srls = Erls, 
   555        calc = [], 
   556        (*asm_thm = [("d2_sqrt_equation1",""),("d2_sqrt_equation1_neg",""),
   557                   ("d2_isolate_div","")],*)
   558        rules = [Thm("d2_isolate_add1",num_str @{d2_isolate_add1),
   559                 (* a+   bx^2=0 -> bx^2=(-1)a*)
   560 		Thm("d2_isolate_add2",num_str @{d2_isolate_add2),
   561                 (* a+    x^2=0 ->  x^2=(-1)a*)
   562 		Thm("d2_sqrt_equation2",num_str @{d2_sqrt_equation2),
   563                 (*  x^2=0 ->    x=0    *)
   564 		Thm("d2_sqrt_equation1",num_str @{d2_sqrt_equation1),
   565                 (* x^2=c   -> x=+-sqrt(c)*)
   566 		Thm("d2_sqrt_equation1_neg",num_str @{d2_sqrt_equation1_neg),
   567                 (* [c<0] x^2=c  -> x=[] *)
   568 		Thm("d2_isolate_div",num_str @{d2_isolate_div)
   569                  (* bx^2=c -> x^2=c/b*)
   570 		],
   571        scr = Script ((term_of o the o (parse thy)) "empty_script")
   572        }:rls);
   573 
   574 (* isolate the bound variable in an d2 equation with pqFormula;
   575    'bdv' is a meta-constant*)
   576 val d2_polyeq_pqFormula_simplify = prep_rls(
   577   Rls {id = "d2_polyeq_pqFormula_simplify", preconds = [],
   578        rew_ord = ("e_rew_ord",e_rew_ord), erls = PolyEq_erls,
   579        srls = Erls, calc = [], 
   580        rules = [Thm("d2_pqformula1",num_str @{d2_pqformula1),
   581                 (* q+px+ x^2=0 *)
   582 		Thm("d2_pqformula1_neg",num_str @{d2_pqformula1_neg),
   583                 (* q+px+ x^2=0 *)
   584 		Thm("d2_pqformula2",num_str @{d2_pqformula2), 
   585                 (* q+px+1x^2=0 *)
   586 		Thm("d2_pqformula2_neg",num_str @{d2_pqformula2_neg),
   587                 (* q+px+1x^2=0 *)
   588 		Thm("d2_pqformula3",num_str @{d2_pqformula3),
   589                 (* q+ x+ x^2=0 *)
   590 		Thm("d2_pqformula3_neg",num_str @{d2_pqformula3_neg), 
   591                 (* q+ x+ x^2=0 *)
   592 		Thm("d2_pqformula4",num_str @{d2_pqformula4),
   593                 (* q+ x+1x^2=0 *)
   594 		Thm("d2_pqformula4_neg",num_str @{d2_pqformula4_neg),
   595                 (* q+ x+1x^2=0 *)
   596 		Thm("d2_pqformula5",num_str @{d2_pqformula5),
   597                 (*   qx+ x^2=0 *)
   598 		Thm("d2_pqformula6",num_str @{d2_pqformula6),
   599                 (*   qx+1x^2=0 *)
   600 		Thm("d2_pqformula7",num_str @{d2_pqformula7),
   601                 (*    x+ x^2=0 *)
   602 		Thm("d2_pqformula8",num_str @{d2_pqformula8),
   603                 (*    x+1x^2=0 *)
   604 		Thm("d2_pqformula9",num_str @{d2_pqformula9),
   605                 (* q   +1x^2=0 *)
   606 		Thm("d2_pqformula9_neg",num_str @{d2_pqformula9_neg),
   607                 (* q   +1x^2=0 *)
   608 		Thm("d2_pqformula10",num_str @{d2_pqformula10),
   609                 (* q   + x^2=0 *)
   610 		Thm("d2_pqformula10_neg",num_str @{d2_pqformula10_neg),
   611                 (* q   + x^2=0 *)
   612 		Thm("d2_sqrt_equation2",num_str @{d2_sqrt_equation2),
   613                 (*       x^2=0 *)
   614 		Thm("d2_sqrt_equation3",num_str @{d2_sqrt_equation3)
   615                (*      1x^2=0 *)
   616 	       ],
   617        scr = Script ((term_of o the o (parse thy)) "empty_script")
   618        }:rls);
   619 
   620 (* isolate the bound variable in an d2 equation with abcFormula; 
   621    'bdv' is a meta-constant*)
   622 val d2_polyeq_abcFormula_simplify = prep_rls(
   623   Rls {id = "d2_polyeq_abcFormula_simplify", preconds = [],
   624        rew_ord = ("e_rew_ord",e_rew_ord), erls = PolyEq_erls,
   625        srls = Erls, calc = [], 
   626        rules = [Thm("d2_abcformula1",num_str @{d2_abcformula1),
   627                 (*c+bx+cx^2=0 *)
   628 		Thm("d2_abcformula1_neg",num_str @{d2_abcformula1_neg),
   629                 (*c+bx+cx^2=0 *)
   630 		Thm("d2_abcformula2",num_str @{d2_abcformula2),
   631                 (*c+ x+cx^2=0 *)
   632 		Thm("d2_abcformula2_neg",num_str @{d2_abcformula2_neg),
   633                 (*c+ x+cx^2=0 *)
   634 		Thm("d2_abcformula3",num_str @{d2_abcformula3), 
   635                 (*c+bx+ x^2=0 *)
   636 		Thm("d2_abcformula3_neg",num_str @{d2_abcformula3_neg),
   637                 (*c+bx+ x^2=0 *)
   638 		Thm("d2_abcformula4",num_str @{d2_abcformula4),
   639                 (*c+ x+ x^2=0 *)
   640 		Thm("d2_abcformula4_neg",num_str @{d2_abcformula4_neg),
   641                 (*c+ x+ x^2=0 *)
   642 		Thm("d2_abcformula5",num_str @{d2_abcformula5),
   643                 (*c+   cx^2=0 *)
   644 		Thm("d2_abcformula5_neg",num_str @{d2_abcformula5_neg),
   645                 (*c+   cx^2=0 *)
   646 		Thm("d2_abcformula6",num_str @{d2_abcformula6),
   647                 (*c+    x^2=0 *)
   648 		Thm("d2_abcformula6_neg",num_str @{d2_abcformula6_neg),
   649                 (*c+    x^2=0 *)
   650 		Thm("d2_abcformula7",num_str @{d2_abcformula7),
   651                 (*  bx+ax^2=0 *)
   652 		Thm("d2_abcformula8",num_str @{d2_abcformula8),
   653                 (*  bx+ x^2=0 *)
   654 		Thm("d2_abcformula9",num_str @{d2_abcformula9),
   655                 (*   x+ax^2=0 *)
   656 		Thm("d2_abcformula10",num_str @{d2_abcformula10),
   657                 (*   x+ x^2=0 *)
   658 		Thm("d2_sqrt_equation2",num_str @{d2_sqrt_equation2),
   659                 (*      x^2=0 *)  
   660 		Thm("d2_sqrt_equation3",num_str @{d2_sqrt_equation3)
   661                (*     bx^2=0 *)  
   662 	       ],
   663        scr = Script ((term_of o the o (parse thy)) "empty_script")
   664        }:rls);
   665 
   666 (* isolate the bound variable in an d2 equation; 
   667    'bdv' is a meta-constant*)
   668 val d2_polyeq_simplify = prep_rls(
   669   Rls {id = "d2_polyeq_simplify", preconds = [],
   670        rew_ord = ("e_rew_ord",e_rew_ord), erls = PolyEq_erls,
   671        srls = Erls, calc = [], 
   672        rules = [Thm("d2_pqformula1",num_str @{d2_pqformula1),
   673                 (* p+qx+ x^2=0 *)
   674 		Thm("d2_pqformula1_neg",num_str @{d2_pqformula1_neg),
   675                 (* p+qx+ x^2=0 *)
   676 		Thm("d2_pqformula2",num_str @{d2_pqformula2),
   677                 (* p+qx+1x^2=0 *)
   678 		Thm("d2_pqformula2_neg",num_str @{d2_pqformula2_neg),
   679                 (* p+qx+1x^2=0 *)
   680 		Thm("d2_pqformula3",num_str @{d2_pqformula3),
   681                 (* p+ x+ x^2=0 *)
   682 		Thm("d2_pqformula3_neg",num_str @{d2_pqformula3_neg),
   683                 (* p+ x+ x^2=0 *)
   684 		Thm("d2_pqformula4",num_str @{d2_pqformula4), 
   685                 (* p+ x+1x^2=0 *)
   686 		Thm("d2_pqformula4_neg",num_str @{d2_pqformula4_neg),
   687                 (* p+ x+1x^2=0 *)
   688 		Thm("d2_abcformula1",num_str @{d2_abcformula1),
   689                 (* c+bx+cx^2=0 *)
   690 		Thm("d2_abcformula1_neg",num_str @{d2_abcformula1_neg),
   691                 (* c+bx+cx^2=0 *)
   692 		Thm("d2_abcformula2",num_str @{d2_abcformula2),
   693                 (* c+ x+cx^2=0 *)
   694 		Thm("d2_abcformula2_neg",num_str @{d2_abcformula2_neg),
   695                 (* c+ x+cx^2=0 *)
   696 		Thm("d2_prescind1",num_str @{d2_prescind1),
   697                 (*   ax+bx^2=0 -> x(a+bx)=0 *)
   698 		Thm("d2_prescind2",num_str @{d2_prescind2),
   699                 (*   ax+ x^2=0 -> x(a+ x)=0 *)
   700 		Thm("d2_prescind3",num_str @{d2_prescind3),
   701                 (*    x+bx^2=0 -> x(1+bx)=0 *)
   702 		Thm("d2_prescind4",num_str @{d2_prescind4),
   703                 (*    x+ x^2=0 -> x(1+ x)=0 *)
   704 		Thm("d2_isolate_add1",num_str @{d2_isolate_add1),
   705                 (* a+   bx^2=0 -> bx^2=(-1)a*)
   706 		Thm("d2_isolate_add2",num_str @{d2_isolate_add2),
   707                 (* a+    x^2=0 ->  x^2=(-1)a*)
   708 		Thm("d2_sqrt_equation1",num_str @{d2_sqrt_equation1),
   709                 (* x^2=c   -> x=+-sqrt(c)*)
   710 		Thm("d2_sqrt_equation1_neg",num_str @{d2_sqrt_equation1_neg),
   711                 (* [c<0] x^2=c   -> x=[]*)
   712 		Thm("d2_sqrt_equation2",num_str @{d2_sqrt_equation2),
   713                 (*  x^2=0 ->    x=0    *)
   714 		Thm("d2_reduce_equation1",num_str @{d2_reduce_equation1),
   715                 (* x(a+bx)=0 -> x=0 | a+bx=0*)
   716 		Thm("d2_reduce_equation2",num_str @{d2_reduce_equation2),
   717                 (* x(a+ x)=0 -> x=0 | a+ x=0*)
   718 		Thm("d2_isolate_div",num_str @{d2_isolate_div)
   719                (* bx^2=c -> x^2=c/b*)
   720 	       ],
   721        scr = Script ((term_of o the o (parse thy)) "empty_script")
   722       }:rls);
   723 
   724 (* -- d3 -- *)
   725 (* isolate the bound variable in an d3 equation; 'bdv' is a meta-constant *)
   726 val d3_polyeq_simplify = prep_rls(
   727   Rls {id = "d3_polyeq_simplify", preconds = [],
   728        rew_ord = ("e_rew_ord",e_rew_ord), erls = PolyEq_erls,
   729        srls = Erls, calc = [], 
   730        rules = 
   731        [Thm("d3_reduce_equation1",num_str @{d3_reduce_equation1),
   732 	(*a*bdv + b*bdv^^^2 + c*bdv^^^3=0) = 
   733         (bdv=0 | (a + b*bdv + c*bdv^^^2=0)*)
   734 	Thm("d3_reduce_equation2",num_str @{d3_reduce_equation2),
   735 	(*  bdv + b*bdv^^^2 + c*bdv^^^3=0) = 
   736         (bdv=0 | (1 + b*bdv + c*bdv^^^2=0)*)
   737 	Thm("d3_reduce_equation3",num_str @{d3_reduce_equation3),
   738 	(*a*bdv +   bdv^^^2 + c*bdv^^^3=0) = 
   739         (bdv=0 | (a +   bdv + c*bdv^^^2=0)*)
   740 	Thm("d3_reduce_equation4",num_str @{d3_reduce_equation4),
   741 	(*  bdv +   bdv^^^2 + c*bdv^^^3=0) = 
   742         (bdv=0 | (1 +   bdv + c*bdv^^^2=0)*)
   743 	Thm("d3_reduce_equation5",num_str @{d3_reduce_equation5),
   744 	(*a*bdv + b*bdv^^^2 +   bdv^^^3=0) = 
   745         (bdv=0 | (a + b*bdv +   bdv^^^2=0)*)
   746 	Thm("d3_reduce_equation6",num_str @{d3_reduce_equation6),
   747 	(*  bdv + b*bdv^^^2 +   bdv^^^3=0) = 
   748         (bdv=0 | (1 + b*bdv +   bdv^^^2=0)*)
   749 	Thm("d3_reduce_equation7",num_str @{d3_reduce_equation7),
   750 	     (*a*bdv +   bdv^^^2 +   bdv^^^3=0) = 
   751              (bdv=0 | (1 +   bdv +   bdv^^^2=0)*)
   752 	Thm("d3_reduce_equation8",num_str @{d3_reduce_equation8),
   753 	     (*  bdv +   bdv^^^2 +   bdv^^^3=0) = 
   754              (bdv=0 | (1 +   bdv +   bdv^^^2=0)*)
   755 	Thm("d3_reduce_equation9",num_str @{d3_reduce_equation9),
   756 	     (*a*bdv             + c*bdv^^^3=0) = 
   757              (bdv=0 | (a         + c*bdv^^^2=0)*)
   758 	Thm("d3_reduce_equation10",num_str @{d3_reduce_equation10),
   759 	     (*  bdv             + c*bdv^^^3=0) = 
   760              (bdv=0 | (1         + c*bdv^^^2=0)*)
   761 	Thm("d3_reduce_equation11",num_str @{d3_reduce_equation11),
   762 	     (*a*bdv             +   bdv^^^3=0) = 
   763              (bdv=0 | (a         +   bdv^^^2=0)*)
   764 	Thm("d3_reduce_equation12",num_str @{d3_reduce_equation12),
   765 	     (*  bdv             +   bdv^^^3=0) = 
   766              (bdv=0 | (1         +   bdv^^^2=0)*)
   767 	Thm("d3_reduce_equation13",num_str @{d3_reduce_equation13),
   768 	     (*        b*bdv^^^2 + c*bdv^^^3=0) = 
   769              (bdv=0 | (    b*bdv + c*bdv^^^2=0)*)
   770 	Thm("d3_reduce_equation14",num_str @{d3_reduce_equation14),
   771 	     (*          bdv^^^2 + c*bdv^^^3=0) = 
   772              (bdv=0 | (      bdv + c*bdv^^^2=0)*)
   773 	Thm("d3_reduce_equation15",num_str @{d3_reduce_equation15),
   774 	     (*        b*bdv^^^2 +   bdv^^^3=0) = 
   775              (bdv=0 | (    b*bdv +   bdv^^^2=0)*)
   776 	Thm("d3_reduce_equation16",num_str @{d3_reduce_equation16),
   777 	     (*          bdv^^^2 +   bdv^^^3=0) = 
   778              (bdv=0 | (      bdv +   bdv^^^2=0)*)
   779 	Thm("d3_isolate_add1",num_str @{d3_isolate_add1),
   780 	     (*[|Not(bdv occurs_in a)|] ==> (a + b*bdv^^^3=0) = 
   781               (bdv=0 | (b*bdv^^^3=a)*)
   782 	Thm("d3_isolate_add2",num_str @{d3_isolate_add2),
   783              (*[|Not(bdv occurs_in a)|] ==> (a +   bdv^^^3=0) = 
   784               (bdv=0 | (  bdv^^^3=a)*)
   785 	Thm("d3_isolate_div",num_str @{d3_isolate_div),
   786         (*[|Not(b=0)|] ==> (b*bdv^^^3=c) = (bdv^^^3=c/b*)
   787         Thm("d3_root_equation2",num_str @{d3_root_equation2),
   788         (*(bdv^^^3=0) = (bdv=0) *)
   789 	Thm("d3_root_equation1",num_str @{d3_root_equation1)
   790        (*bdv^^^3=c) = (bdv = nroot 3 c*)
   791        ],
   792        scr = Script ((term_of o the o (parse thy)) "empty_script")
   793       }:rls);
   794 
   795 (* -- d4 -- *)
   796 (*isolate the bound variable in an d4 equation; 'bdv' is a meta-constant*)
   797 val d4_polyeq_simplify = prep_rls(
   798   Rls {id = "d4_polyeq_simplify", preconds = [],
   799        rew_ord = ("e_rew_ord",e_rew_ord), erls = PolyEq_erls,
   800        srls = Erls, calc = [], 
   801        rules = 
   802        [Thm("d4_sub_u1",num_str @{d4_sub_u1)  
   803        (* ax^4+bx^2+c=0 -> x=+-sqrt(ax^2+bx^+c) *)
   804        ],
   805        scr = Script ((term_of o the o (parse thy)) "empty_script")
   806       }:rls);
   807   
   808 ruleset' := 
   809 overwritelthy @{theory} 
   810               (!ruleset',
   811                [("d0_polyeq_simplify", d0_polyeq_simplify),
   812                 ("d1_polyeq_simplify", d1_polyeq_simplify),
   813                 ("d2_polyeq_simplify", d2_polyeq_simplify),
   814                 ("d2_polyeq_bdv_only_simplify", d2_polyeq_bdv_only_simplify),
   815                 ("d2_polyeq_sq_only_simplify", d2_polyeq_sq_only_simplify),
   816                 ("d2_polyeq_pqFormula_simplify", d2_polyeq_pqFormula_simplify),
   817                 ("d2_polyeq_abcFormula_simplify", 
   818                  d2_polyeq_abcFormula_simplify),
   819                 ("d3_polyeq_simplify", d3_polyeq_simplify),
   820 		("d4_polyeq_simplify", d4_polyeq_simplify)
   821 	      ]);
   822     
   823 (*------------------------problems------------------------*)
   824 (*
   825 (get_pbt ["degree_2","polynomial","univariate","equation"]);
   826 show_ptyps(); 
   827 *)
   828 
   829 (*-------------------------poly-----------------------*)
   830 store_pbt
   831  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly" [] e_pblID
   832  (["polynomial","univariate","equation"],
   833   [("#Given" ,["equality e_","solveFor v_"]),
   834    ("#Where" ,["~((e_::bool) is_ratequation_in (v_::real))",
   835 	       "~((lhs e_) is_rootTerm_in (v_::real))",
   836 	       "~((rhs e_) is_rootTerm_in (v_::real))"]),
   837    ("#Find"  ,["solutions v_i_"])
   838    ],
   839   PolyEq_prls, SOME "solve (e_::bool, v_)",
   840   []));
   841 (*--- d0 ---*)
   842 store_pbt
   843  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg0" [] e_pblID
   844  (["degree_0","polynomial","univariate","equation"],
   845   [("#Given" ,["equality e_","solveFor v_"]),
   846    ("#Where" ,["matches (?a = 0) e_",
   847 	       "(lhs e_) is_poly_in v_",
   848 	       "((lhs e_) has_degree_in v_ ) = 0"
   849 	      ]),
   850    ("#Find"  ,["solutions v_i_"])
   851   ],
   852   PolyEq_prls, SOME "solve (e_::bool, v_)",
   853   [["PolyEq","solve_d0_polyeq_equation"]]));
   854 
   855 (*--- d1 ---*)
   856 store_pbt
   857  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg1" [] e_pblID
   858  (["degree_1","polynomial","univariate","equation"],
   859   [("#Given" ,["equality e_","solveFor v_"]),
   860    ("#Where" ,["matches (?a = 0) e_",
   861 	       "(lhs e_) is_poly_in v_",
   862 	       "((lhs e_) has_degree_in v_ ) = 1"
   863 	      ]),
   864    ("#Find"  ,["solutions v_i_"])
   865   ],
   866   PolyEq_prls, SOME "solve (e_::bool, v_)",
   867   [["PolyEq","solve_d1_polyeq_equation"]]));
   868 
   869 (*--- d2 ---*)
   870 store_pbt
   871  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg2" [] e_pblID
   872  (["degree_2","polynomial","univariate","equation"],
   873   [("#Given" ,["equality e_","solveFor v_"]),
   874    ("#Where" ,["matches (?a = 0) e_",
   875 	       "(lhs e_) is_poly_in v_ ",
   876 	       "((lhs e_) has_degree_in v_ ) = 2"]),
   877    ("#Find"  ,["solutions v_i_"])
   878   ],
   879   PolyEq_prls, SOME "solve (e_::bool, v_)",
   880   [["PolyEq","solve_d2_polyeq_equation"]]));
   881 
   882  store_pbt
   883  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg2_sqonly" [] e_pblID
   884  (["sq_only","degree_2","polynomial","univariate","equation"],
   885   [("#Given" ,["equality e_","solveFor v_"]),
   886    ("#Where" ,["matches ( ?a +    ?v_^^^2 = 0) e_ | " ^
   887 	       "matches ( ?a + ?b*?v_^^^2 = 0) e_ | " ^
   888 	       "matches (         ?v_^^^2 = 0) e_ | " ^
   889 	       "matches (      ?b*?v_^^^2 = 0) e_" ,
   890 	       "Not (matches (?a +    ?v_ +    ?v_^^^2 = 0) e_) &" ^
   891 	       "Not (matches (?a + ?b*?v_ +    ?v_^^^2 = 0) e_) &" ^
   892 	       "Not (matches (?a +    ?v_ + ?c*?v_^^^2 = 0) e_) &" ^
   893 	       "Not (matches (?a + ?b*?v_ + ?c*?v_^^^2 = 0) e_) &" ^
   894 	       "Not (matches (        ?v_ +    ?v_^^^2 = 0) e_) &" ^
   895 	       "Not (matches (     ?b*?v_ +    ?v_^^^2 = 0) e_) &" ^
   896 	       "Not (matches (        ?v_ + ?c*?v_^^^2 = 0) e_) &" ^
   897 	       "Not (matches (     ?b*?v_ + ?c*?v_^^^2 = 0) e_)"]),
   898    ("#Find"  ,["solutions v_i_"])
   899   ],
   900   PolyEq_prls, SOME "solve (e_::bool, v_)",
   901   [["PolyEq","solve_d2_polyeq_sqonly_equation"]]));
   902 
   903 store_pbt
   904  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg2_bdvonly" [] e_pblID
   905  (["bdv_only","degree_2","polynomial","univariate","equation"],
   906   [("#Given" ,["equality e_","solveFor v_"]),
   907    ("#Where" ,["matches (?a*?v_ +    ?v_^^^2 = 0) e_ | " ^
   908 	       "matches (   ?v_ +    ?v_^^^2 = 0) e_ | " ^
   909 	       "matches (   ?v_ + ?b*?v_^^^2 = 0) e_ | " ^
   910 	       "matches (?a*?v_ + ?b*?v_^^^2 = 0) e_ | " ^
   911 	       "matches (            ?v_^^^2 = 0) e_ | " ^
   912 	       "matches (         ?b*?v_^^^2 = 0) e_ "]),
   913    ("#Find"  ,["solutions v_i_"])
   914   ],
   915   PolyEq_prls, SOME "solve (e_::bool, v_)",
   916   [["PolyEq","solve_d2_polyeq_bdvonly_equation"]]));
   917 
   918 store_pbt
   919  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg2_pq" [] e_pblID
   920  (["pqFormula","degree_2","polynomial","univariate","equation"],
   921   [("#Given" ,["equality e_","solveFor v_"]),
   922    ("#Where" ,["matches (?a + 1*?v_^^^2 = 0) e_ | " ^
   923 	       "matches (?a +   ?v_^^^2 = 0) e_"]),
   924    ("#Find"  ,["solutions v_i_"])
   925   ],
   926   PolyEq_prls, SOME "solve (e_::bool, v_)",
   927   [["PolyEq","solve_d2_polyeq_pq_equation"]]));
   928 
   929 store_pbt
   930  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg2_abc" [] e_pblID
   931  (["abcFormula","degree_2","polynomial","univariate","equation"],
   932   [("#Given" ,["equality e_","solveFor v_"]),
   933    ("#Where" ,["matches (?a +    ?v_^^^2 = 0) e_ | " ^
   934 	       "matches (?a + ?b*?v_^^^2 = 0) e_"]),
   935    ("#Find"  ,["solutions v_i_"])
   936   ],
   937   PolyEq_prls, SOME "solve (e_::bool, v_)",
   938   [["PolyEq","solve_d2_polyeq_abc_equation"]]));
   939 
   940 (*--- d3 ---*)
   941 store_pbt
   942  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg3" [] e_pblID
   943  (["degree_3","polynomial","univariate","equation"],
   944   [("#Given" ,["equality e_","solveFor v_"]),
   945    ("#Where" ,["matches (?a = 0) e_",
   946 	       "(lhs e_) is_poly_in v_ ",
   947 	       "((lhs e_) has_degree_in v_) = 3"]),
   948    ("#Find"  ,["solutions v_i_"])
   949   ],
   950   PolyEq_prls, SOME "solve (e_::bool, v_)",
   951   [["PolyEq","solve_d3_polyeq_equation"]]));
   952 
   953 (*--- d4 ---*)
   954 store_pbt
   955  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_deg4" [] e_pblID
   956  (["degree_4","polynomial","univariate","equation"],
   957   [("#Given" ,["equality e_","solveFor v_"]),
   958    ("#Where" ,["matches (?a = 0) e_",
   959 	       "(lhs e_) is_poly_in v_ ",
   960 	       "((lhs e_) has_degree_in v_) = 4"]),
   961    ("#Find"  ,["solutions v_i_"])
   962   ],
   963   PolyEq_prls, SOME "solve (e_::bool, v_)",
   964   [(*["PolyEq","solve_d4_polyeq_equation"]*)]));
   965 
   966 (*--- normalize ---*)
   967 store_pbt
   968  (prep_pbt (theory "PolyEq") "pbl_equ_univ_poly_norm" [] e_pblID
   969  (["normalize","polynomial","univariate","equation"],
   970   [("#Given" ,["equality e_","solveFor v_"]),
   971    ("#Where" ,["(Not((matches (?a = 0 ) e_ ))) |" ^
   972 	       "(Not(((lhs e_) is_poly_in v_)))"]),
   973    ("#Find"  ,["solutions v_i_"])
   974   ],
   975   PolyEq_prls, SOME "solve (e_::bool, v_)",
   976   [["PolyEq","normalize_poly"]]));
   977 (*-------------------------expanded-----------------------*)
   978 store_pbt
   979  (prep_pbt (theory "PolyEq") "pbl_equ_univ_expand" [] e_pblID
   980  (["expanded","univariate","equation"],
   981   [("#Given" ,["equality e_","solveFor v_"]),
   982    ("#Where" ,["matches (?a = 0) e_",
   983 	       "(lhs e_) is_expanded_in v_ "]),
   984    ("#Find"  ,["solutions v_i_"])
   985    ],
   986   PolyEq_prls, SOME "solve (e_::bool, v_)",
   987   []));
   988 
   989 (*--- d2 ---*)
   990 store_pbt
   991  (prep_pbt (theory "PolyEq") "pbl_equ_univ_expand_deg2" [] e_pblID
   992  (["degree_2","expanded","univariate","equation"],
   993   [("#Given" ,["equality e_","solveFor v_"]),
   994    ("#Where" ,["((lhs e_) has_degree_in v_) = 2"]),
   995    ("#Find"  ,["solutions v_i_"])
   996   ],
   997   PolyEq_prls, SOME "solve (e_::bool, v_)",
   998   [["PolyEq","complete_square"]]));
   999 
  1000 
  1001 "-------------------------methods-----------------------";
  1002 store_met
  1003  (prep_met (theory "PolyEq") "met_polyeq" [] e_metID
  1004  (["PolyEq"],
  1005    [],
  1006    {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = e_rls, prls=e_rls,
  1007     crls=PolyEq_crls, nrls=norm_Rational}, "empty_script"));
  1008 
  1009 store_met
  1010  (prep_met (theory "PolyEq") "met_polyeq_norm" [] e_metID
  1011  (["PolyEq","normalize_poly"],
  1012    [("#Given" ,["equality e_","solveFor v_"]),
  1013    ("#Where" ,["(Not((matches (?a = 0 ) e_ ))) |" ^
  1014 	       "(Not(((lhs e_) is_poly_in v_)))"]),
  1015    ("#Find"  ,["solutions v_i_"])
  1016   ],
  1017    {rew_ord'="termlessI",
  1018     rls'=PolyEq_erls,
  1019     srls=e_rls,
  1020     prls=PolyEq_prls,
  1021     calc=[],
  1022     crls=PolyEq_crls, nrls=norm_Rational
  1023     "Script Normalize_poly (e_::bool) (v_::real) =                     " ^
  1024     "(let e_ =((Try         (Rewrite     all_left          False)) @@  " ^ 
  1025     "          (Try (Repeat (Rewrite     makex1_x         False))) @@  " ^ 
  1026     "          (Try (Repeat (Rewrite_Set expand_binoms    False))) @@  " ^ 
  1027     "          (Try (Repeat (Rewrite_Set_Inst [(bdv,v_::real)]         " ^
  1028     "                                 make_ratpoly_in     False))) @@  " ^
  1029     "          (Try (Repeat (Rewrite_Set polyeq_simplify  False)))) e_ " ^
  1030     " in (SubProblem (PolyEq_,[polynomial,univariate,equation],        " ^
  1031     "                [no_met]) [bool_ e_, real_ v_]))"
  1032    ));
  1033 
  1034 store_met
  1035  (prep_met (theory "PolyEq") "met_polyeq_d0" [] e_metID
  1036  (["PolyEq","solve_d0_polyeq_equation"],
  1037    [("#Given" ,["equality e_","solveFor v_"]),
  1038    ("#Where" ,["(lhs e_) is_poly_in v_ ",
  1039 	       "((lhs e_) has_degree_in v_) = 0"]),
  1040    ("#Find"  ,["solutions v_i_"])
  1041   ],
  1042    {rew_ord'="termlessI",
  1043     rls'=PolyEq_erls,
  1044     srls=e_rls,
  1045     prls=PolyEq_prls,
  1046     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1047     crls=PolyEq_crls, nrls=norm_Rational},
  1048    "Script Solve_d0_polyeq_equation  (e_::bool) (v_::real)  = " ^
  1049     "(let e_ =  ((Try (Rewrite_Set_Inst [(bdv,v_::real)]      " ^
  1050     "                  d0_polyeq_simplify  False))) e_        " ^
  1051     " in ((Or_to_List e_)::bool list))"
  1052  ));
  1053 
  1054 store_met
  1055  (prep_met (theory "PolyEq") "met_polyeq_d1" [] e_metID
  1056  (["PolyEq","solve_d1_polyeq_equation"],
  1057    [("#Given" ,["equality e_","solveFor v_"]),
  1058    ("#Where" ,["(lhs e_) is_poly_in v_ ",
  1059 	       "((lhs e_) has_degree_in v_) = 1"]),
  1060    ("#Find"  ,["solutions v_i_"])
  1061   ],
  1062    {rew_ord'="termlessI",
  1063     rls'=PolyEq_erls,
  1064     srls=e_rls,
  1065     prls=PolyEq_prls,
  1066     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1067     crls=PolyEq_crls, nrls=norm_Rational(*,
  1068     (*    asm_rls=["d1_polyeq_simplify"],*)
  1069     asm_rls=[],
  1070     asm_thm=[("d1_isolate_div","")]*)},
  1071    "Script Solve_d1_polyeq_equation  (e_::bool) (v_::real)  =   " ^
  1072     "(let e_ =  ((Try (Rewrite_Set_Inst [(bdv,v_::real)]        " ^
  1073     "                  d1_polyeq_simplify   True))          @@  " ^
  1074     "            (Try (Rewrite_Set polyeq_simplify  False)) @@  " ^
  1075     "            (Try (Rewrite_Set norm_Rational_parenthesized False))) e_;" ^
  1076     " (L_::bool list) = ((Or_to_List e_)::bool list)            " ^
  1077     " in Check_elementwise L_ {(v_::real). Assumptions} )"
  1078  ));
  1079 
  1080 store_met
  1081  (prep_met (theory "PolyEq") "met_polyeq_d22" [] e_metID
  1082  (["PolyEq","solve_d2_polyeq_equation"],
  1083    [("#Given" ,["equality e_","solveFor v_"]),
  1084    ("#Where" ,["(lhs e_) is_poly_in v_ ",
  1085 	       "((lhs e_) has_degree_in v_) = 2"]),
  1086    ("#Find"  ,["solutions v_i_"])
  1087   ],
  1088    {rew_ord'="termlessI",
  1089     rls'=PolyEq_erls,
  1090     srls=e_rls,
  1091     prls=PolyEq_prls,
  1092     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1093     crls=PolyEq_crls, nrls=norm_Rational},
  1094    "Script Solve_d2_polyeq_equation  (e_::bool) (v_::real) =      " ^
  1095     "  (let e_ = ((Try (Rewrite_Set_Inst [(bdv,v_::real)]         " ^
  1096     "                    d2_polyeq_simplify           True)) @@   " ^
  1097     "             (Try (Rewrite_Set polyeq_simplify   False)) @@  " ^
  1098     "             (Try (Rewrite_Set_Inst [(bdv,v_::real)]         " ^
  1099     "                    d1_polyeq_simplify            True)) @@  " ^
  1100     "            (Try (Rewrite_Set polyeq_simplify    False)) @@  " ^
  1101     "            (Try (Rewrite_Set norm_Rational_parenthesized False))) e_;" ^
  1102     " (L_::bool list) = ((Or_to_List e_)::bool list)              " ^
  1103     " in Check_elementwise L_ {(v_::real). Assumptions} )"
  1104  ));
  1105 
  1106 store_met
  1107  (prep_met (theory "PolyEq") "met_polyeq_d2_bdvonly" [] e_metID
  1108  (["PolyEq","solve_d2_polyeq_bdvonly_equation"],
  1109    [("#Given" ,["equality e_","solveFor v_"]),
  1110    ("#Where" ,["(lhs e_) is_poly_in v_ ",
  1111 	       "((lhs e_) has_degree_in v_) = 2"]),
  1112    ("#Find"  ,["solutions v_i_"])
  1113   ],
  1114    {rew_ord'="termlessI",
  1115     rls'=PolyEq_erls,
  1116     srls=e_rls,
  1117     prls=PolyEq_prls,
  1118     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1119     crls=PolyEq_crls, nrls=norm_Rational},
  1120    "Script Solve_d2_polyeq_bdvonly_equation  (e_::bool) (v_::real) =" ^
  1121     "  (let e_ = ((Try (Rewrite_Set_Inst [(bdv,v_::real)]         " ^
  1122     "                   d2_polyeq_bdv_only_simplify    True)) @@  " ^
  1123     "             (Try (Rewrite_Set polyeq_simplify   False)) @@  " ^
  1124     "             (Try (Rewrite_Set_Inst [(bdv,v_::real)]         " ^
  1125     "                   d1_polyeq_simplify             True)) @@  " ^
  1126     "            (Try (Rewrite_Set polyeq_simplify    False)) @@  " ^
  1127     "            (Try (Rewrite_Set norm_Rational_parenthesized False))) e_;" ^
  1128     " (L_::bool list) = ((Or_to_List e_)::bool list)              " ^
  1129     " in Check_elementwise L_ {(v_::real). Assumptions} )"
  1130  ));
  1131 
  1132 store_met
  1133  (prep_met (theory "PolyEq") "met_polyeq_d2_sqonly" [] e_metID
  1134  (["PolyEq","solve_d2_polyeq_sqonly_equation"],
  1135    [("#Given" ,["equality e_","solveFor v_"]),
  1136    ("#Where" ,["(lhs e_) is_poly_in v_ ",
  1137 	       "((lhs e_) has_degree_in v_) = 2"]),
  1138    ("#Find"  ,["solutions v_i_"])
  1139   ],
  1140    {rew_ord'="termlessI",
  1141     rls'=PolyEq_erls,
  1142     srls=e_rls,
  1143     prls=PolyEq_prls,
  1144     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1145     crls=PolyEq_crls, nrls=norm_Rational},
  1146    "Script Solve_d2_polyeq_sqonly_equation  (e_::bool) (v_::real) =" ^
  1147     "  (let e_ = ((Try (Rewrite_Set_Inst [(bdv,v_::real)]          " ^
  1148     "                   d2_polyeq_sq_only_simplify     True)) @@   " ^
  1149     "            (Try (Rewrite_Set polyeq_simplify    False)) @@   " ^
  1150     "            (Try (Rewrite_Set norm_Rational_parenthesized False))) e_; " ^
  1151     " (L_::bool list) = ((Or_to_List e_)::bool list)               " ^
  1152     " in Check_elementwise L_ {(v_::real). Assumptions} )"
  1153  ));
  1154 
  1155 store_met
  1156  (prep_met (theory "PolyEq") "met_polyeq_d2_pq" [] e_metID
  1157  (["PolyEq","solve_d2_polyeq_pq_equation"],
  1158    [("#Given" ,["equality e_","solveFor v_"]),
  1159    ("#Where" ,["(lhs e_) is_poly_in v_ ",
  1160 	       "((lhs e_) has_degree_in v_) = 2"]),
  1161    ("#Find"  ,["solutions v_i_"])
  1162   ],
  1163    {rew_ord'="termlessI",
  1164     rls'=PolyEq_erls,
  1165     srls=e_rls,
  1166     prls=PolyEq_prls,
  1167     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1168     crls=PolyEq_crls, nrls=norm_Rational},
  1169    "Script Solve_d2_polyeq_pq_equation  (e_::bool) (v_::real) =   " ^
  1170     "  (let e_ = ((Try (Rewrite_Set_Inst [(bdv,v_::real)]         " ^
  1171     "                   d2_polyeq_pqFormula_simplify   True)) @@  " ^
  1172     "            (Try (Rewrite_Set polyeq_simplify    False)) @@  " ^
  1173     "            (Try (Rewrite_Set norm_Rational_parenthesized False))) e_;" ^
  1174     " (L_::bool list) = ((Or_to_List e_)::bool list)              " ^
  1175     " in Check_elementwise L_ {(v_::real). Assumptions} )"
  1176  ));
  1177 
  1178 store_met
  1179  (prep_met (theory "PolyEq") "met_polyeq_d2_abc" [] e_metID
  1180  (["PolyEq","solve_d2_polyeq_abc_equation"],
  1181    [("#Given" ,["equality e_","solveFor v_"]),
  1182    ("#Where" ,["(lhs e_) is_poly_in v_ ",
  1183 	       "((lhs e_) has_degree_in v_) = 2"]),
  1184    ("#Find"  ,["solutions v_i_"])
  1185   ],
  1186    {rew_ord'="termlessI",
  1187     rls'=PolyEq_erls,
  1188     srls=e_rls,
  1189     prls=PolyEq_prls,
  1190     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1191     crls=PolyEq_crls, nrls=norm_Rational},
  1192    "Script Solve_d2_polyeq_abc_equation  (e_::bool) (v_::real) =   " ^
  1193     "  (let e_ = ((Try (Rewrite_Set_Inst [(bdv,v_::real)]          " ^
  1194     "                   d2_polyeq_abcFormula_simplify   True)) @@  " ^
  1195     "            (Try (Rewrite_Set polyeq_simplify     False)) @@  " ^
  1196     "            (Try (Rewrite_Set norm_Rational_parenthesized False))) e_;" ^
  1197     " (L_::bool list) = ((Or_to_List e_)::bool list)               " ^
  1198     " in Check_elementwise L_ {(v_::real). Assumptions} )"
  1199  ));
  1200 
  1201 store_met
  1202  (prep_met (theory "PolyEq") "met_polyeq_d3" [] e_metID
  1203  (["PolyEq","solve_d3_polyeq_equation"],
  1204    [("#Given" ,["equality e_","solveFor v_"]),
  1205    ("#Where" ,["(lhs e_) is_poly_in v_ ",
  1206 	       "((lhs e_) has_degree_in v_) = 3"]),
  1207    ("#Find"  ,["solutions v_i_"])
  1208   ],
  1209    {rew_ord'="termlessI",
  1210     rls'=PolyEq_erls,
  1211     srls=e_rls,
  1212     prls=PolyEq_prls,
  1213     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1214     crls=PolyEq_crls, nrls=norm_Rational},
  1215    "Script Solve_d3_polyeq_equation  (e_::bool) (v_::real) =     " ^
  1216     "  (let e_ = ((Try (Rewrite_Set_Inst [(bdv,v_::real)]        " ^
  1217     "                    d3_polyeq_simplify           True)) @@  " ^
  1218     "             (Try (Rewrite_Set polyeq_simplify  False)) @@  " ^
  1219     "             (Try (Rewrite_Set_Inst [(bdv,v_::real)]        " ^
  1220     "                    d2_polyeq_simplify           True)) @@  " ^
  1221     "             (Try (Rewrite_Set polyeq_simplify  False)) @@  " ^
  1222     "             (Try (Rewrite_Set_Inst [(bdv,v_::real)]        " ^   
  1223     "                    d1_polyeq_simplify           True)) @@  " ^
  1224     "             (Try (Rewrite_Set polyeq_simplify  False)) @@  " ^
  1225     "             (Try (Rewrite_Set norm_Rational_parenthesized False))) e_;" ^
  1226     " (L_::bool list) = ((Or_to_List e_)::bool list)             " ^
  1227     " in Check_elementwise L_ {(v_::real). Assumptions} )"
  1228    ));
  1229 
  1230  (*.solves all expanded (ie. normalized) terms of degree 2.*) 
  1231  (*Oct.02 restriction: 'eval_true 0 =< discriminant' ony for integer values
  1232    by 'PolyEq_erls'; restricted until Float.thy is implemented*)
  1233 store_met
  1234  (prep_met (theory "PolyEq") "met_polyeq_complsq" [] e_metID
  1235  (["PolyEq","complete_square"],
  1236    [("#Given" ,["equality e_","solveFor v_"]),
  1237    ("#Where" ,["matches (?a = 0) e_", 
  1238 	       "((lhs e_) has_degree_in v_) = 2"]),
  1239    ("#Find"  ,["solutions v_i_"])
  1240   ],
  1241    {rew_ord'="termlessI",rls'=PolyEq_erls,srls=e_rls,prls=PolyEq_prls,
  1242     calc=[("sqrt", ("Root.sqrt", eval_sqrt "#sqrt_"))],
  1243     crls=PolyEq_crls, nrls=norm_Rational},
  1244    "Script Complete_square (e_::bool) (v_::real) =                          " ^
  1245    "(let e_ = ((Try (Rewrite_Set_Inst [(bdv,v_)] cancel_leading_coeff True))" ^
  1246    "        @@ (Try (Rewrite_Set_Inst [(bdv,v_)] complete_square True))     " ^
  1247    "        @@ (Try (Rewrite square_explicit1 False))                       " ^
  1248    "        @@ (Try (Rewrite square_explicit2 False))                       " ^
  1249    "        @@ (Rewrite root_plus_minus True)                               " ^
  1250    "        @@ (Try (Repeat (Rewrite_Inst [(bdv,v_)] bdv_explicit1 False))) " ^
  1251    "        @@ (Try (Repeat (Rewrite_Inst [(bdv,v_)] bdv_explicit2 False))) " ^
  1252    "        @@ (Try (Repeat                                                 " ^
  1253    "                  (Rewrite_Inst [(bdv,v_)] bdv_explicit3 False)))       " ^
  1254    "        @@ (Try (Rewrite_Set calculate_RootRat False))                  " ^
  1255    "        @@ (Try (Repeat (Calculate sqrt_)))) e_                         " ^
  1256    " in ((Or_to_List e_)::bool list))"
  1257    ));
  1258 
  1259 
  1260 (* termorder hacked by MG *)
  1261 local (*. for make_polynomial_in .*)
  1262 
  1263 open Term;  (* for type order = EQUAL | LESS | GREATER *)
  1264 
  1265 fun pr_ord EQUAL = "EQUAL"
  1266   | pr_ord LESS  = "LESS"
  1267   | pr_ord GREATER = "GREATER";
  1268 
  1269 fun dest_hd' x (Const (a, T)) = (((a, 0), T), 0)
  1270   | dest_hd' x (t as Free (a, T)) =
  1271     if x = t then ((("|||||||||||||", 0), T), 0)                        (*WN*)
  1272     else (((a, 0), T), 1)
  1273   | dest_hd' x (Var v) = (v, 2)
  1274   | dest_hd' x (Bound i) = ((("", i), dummyT), 3)
  1275   | dest_hd' x (Abs (_, T, _)) = ((("", 0), T), 4);
  1276 
  1277 fun size_of_term' x (Const ("Atools.pow",_) $ Free (var,_) $ Free (pot,_)) =
  1278     (case x of                                                          (*WN*)
  1279 	    (Free (xstr,_)) => 
  1280 		(if xstr = var then 1000*(the (int_of_str pot)) else 3)
  1281 	  | _ => raise error ("size_of_term' called with subst = "^
  1282 			      (term2str x)))
  1283   | size_of_term' x (Free (subst,_)) =
  1284     (case x of
  1285 	    (Free (xstr,_)) => (if xstr = subst then 1000 else 1)
  1286 	  | _ => raise error ("size_of_term' called with subst = "^
  1287 			  (term2str x)))
  1288   | size_of_term' x (Abs (_,_,body)) = 1 + size_of_term' x body
  1289   | size_of_term' x (f$t) = size_of_term' x f  +  size_of_term' x t
  1290   | size_of_term' x _ = 1;
  1291 
  1292 
  1293 fun term_ord' x pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
  1294       (case term_ord' x pr thy (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
  1295   | term_ord' x pr thy (t, u) =
  1296       (if pr then 
  1297 	 let
  1298 	   val (f, ts) = strip_comb t and (g, us) = strip_comb u;
  1299 	   val _=writeln("t= f@ts= \""^
  1300 	      ((Syntax.string_of_term (thy2ctxt thy)) f)^"\" @ \"["^
  1301 	      (commas(map(string_of_cterm o cterm_of(sign_of thy)) ts))^"]\"");
  1302 	   val _=writeln("u= g@us= \""^
  1303 	      ((Syntax.string_of_term (thy2ctxt thy)) g)^"\" @ \"["^
  1304 	      (commas(map(string_of_cterm o cterm_of(sign_of thy)) us))^"]\"");
  1305 	   val _=writeln("size_of_term(t,u)= ("^
  1306 	      (string_of_int(size_of_term' x t))^", "^
  1307 	      (string_of_int(size_of_term' x u))^")");
  1308 	   val _=writeln("hd_ord(f,g)      = "^((pr_ord o (hd_ord x))(f,g)));
  1309 	   val _=writeln("terms_ord(ts,us) = "^
  1310 			   ((pr_ord o (terms_ord x) str false)(ts,us)));
  1311 	   val _=writeln("-------");
  1312 	 in () end
  1313        else ();
  1314 	 case int_ord (size_of_term' x t, size_of_term' x u) of
  1315 	   EQUAL =>
  1316 	     let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
  1317 	       (case hd_ord x (f, g) of EQUAL => (terms_ord x str pr) (ts, us) 
  1318 	     | ord => ord)
  1319 	     end
  1320 	 | ord => ord)
  1321 and hd_ord x (f, g) =                                        (* ~ term.ML *)
  1322   prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd' x f, 
  1323 						     dest_hd' x g)
  1324 and terms_ord x str pr (ts, us) = 
  1325     list_ord (term_ord' x pr (assoc_thy "Isac.thy"))(ts, us);
  1326 in
  1327 
  1328 fun ord_make_polynomial_in (pr:bool) thy subst tu = 
  1329     let
  1330 	(* val _=writeln("*** subs variable is: "^(subst2str subst)); *)
  1331     in
  1332 	case subst of
  1333 	    (_,x)::_ => (term_ord' x pr thy tu = LESS)
  1334 	  | _ => raise error ("ord_make_polynomial_in called with subst = "^
  1335 			  (subst2str subst))
  1336     end;
  1337 end;
  1338 
  1339 val order_add_mult_in = prep_rls(
  1340   Rls{id = "order_add_mult_in", preconds = [], 
  1341       rew_ord = ("ord_make_polynomial_in",
  1342 		 ord_make_polynomial_in false Poly.thy),
  1343       erls = e_rls,srls = Erls,
  1344       calc = [],
  1345       (*asm_thm = [],*)
  1346       rules = [Thm ("real_mult_commute",num_str @{real_mult_commute),
  1347 	       (* z * w = w * z *)
  1348 	       Thm ("real_mult_left_commute",num_str @{real_mult_left_commute),
  1349 	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
  1350 	       Thm ("real_mult_assoc",num_str @{real_mult_assoc),		
  1351 	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
  1352 	       Thm ("add_commute",num_str @{thm add_commute}),	
  1353 	       (*z + w = w + z*)
  1354 	       Thm ("add_left_commute",num_str @{thm add_left_commute}),
  1355 	       (*x + (y + z) = y + (x + z)*)
  1356 	       Thm ("add_assoc",num_str @{thm add_assoc})	               
  1357 	       (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
  1358 	       ], scr = EmptyScr}:rls);
  1359 
  1360 val collect_bdv = prep_rls(
  1361   Rls{id = "collect_bdv", preconds = [], 
  1362       rew_ord = ("dummy_ord", dummy_ord),
  1363       erls = e_rls,srls = Erls,
  1364       calc = [],
  1365       (*asm_thm = [],*)
  1366       rules = [Thm ("bdv_collect_1",num_str @{bdv_collect_1),
  1367 	       Thm ("bdv_collect_2",num_str @{bdv_collect_2),
  1368 	       Thm ("bdv_collect_3",num_str @{bdv_collect_3),
  1369 
  1370 	       Thm ("bdv_collect_assoc1_1",num_str @{bdv_collect_assoc1_1),
  1371 	       Thm ("bdv_collect_assoc1_2",num_str @{bdv_collect_assoc1_2),
  1372 	       Thm ("bdv_collect_assoc1_3",num_str @{bdv_collect_assoc1_3),
  1373 
  1374 	       Thm ("bdv_collect_assoc2_1",num_str @{bdv_collect_assoc2_1),
  1375 	       Thm ("bdv_collect_assoc2_2",num_str @{bdv_collect_assoc2_2),
  1376 	       Thm ("bdv_collect_assoc2_3",num_str @{bdv_collect_assoc2_3),
  1377 
  1378 
  1379 	       Thm ("bdv_n_collect_1",num_str @{bdv_n_collect_1),
  1380 	       Thm ("bdv_n_collect_2",num_str @{bdv_n_collect_2),
  1381 	       Thm ("bdv_n_collect_3",num_str @{bdv_n_collect_3),
  1382 
  1383 	       Thm ("bdv_n_collect_assoc1_1",num_str @{bdv_n_collect_assoc1_1),
  1384 	       Thm ("bdv_n_collect_assoc1_2",num_str @{bdv_n_collect_assoc1_2),
  1385 	       Thm ("bdv_n_collect_assoc1_3",num_str @{bdv_n_collect_assoc1_3),
  1386 
  1387 	       Thm ("bdv_n_collect_assoc2_1",num_str @{bdv_n_collect_assoc2_1),
  1388 	       Thm ("bdv_n_collect_assoc2_2",num_str @{bdv_n_collect_assoc2_2),
  1389 	       Thm ("bdv_n_collect_assoc2_3",num_str @{bdv_n_collect_assoc2_3)
  1390 	       ], scr = EmptyScr}:rls);
  1391 
  1392 (*.transforms an arbitrary term without roots to a polynomial [4] 
  1393    according to knowledge/Poly.sml.*) 
  1394 val make_polynomial_in = prep_rls(
  1395   Seq {id = "make_polynomial_in", preconds = []:term list, 
  1396        rew_ord = ("dummy_ord", dummy_ord),
  1397       erls = Atools_erls, srls = Erls,
  1398       calc = [], (*asm_thm = [],*)
  1399       rules = [Rls_ expand_poly,
  1400 	       Rls_ order_add_mult_in,
  1401 	       Rls_ simplify_power,
  1402 	       Rls_ collect_numerals,
  1403 	       Rls_ reduce_012,
  1404 	       Thm ("realpow_oneI",num_str @{realpow_oneI),
  1405 	       Rls_ discard_parentheses,
  1406 	       Rls_ collect_bdv
  1407 	       ],
  1408       scr = EmptyScr
  1409       }:rls);     
  1410 
  1411 val separate_bdvs = 
  1412     append_rls "separate_bdvs"
  1413 	       collect_bdv
  1414 	       [Thm ("separate_bdv", num_str @{separate_bdv),
  1415 		(*"?a * ?bdv / ?b = ?a / ?b * ?bdv"*)
  1416 		Thm ("separate_bdv_n", num_str @{separate_bdv_n),
  1417 		Thm ("separate_1_bdv", num_str @{separate_1_bdv),
  1418 		(*"?bdv / ?b = (1 / ?b) * ?bdv"*)
  1419 		Thm ("separate_1_bdv_n", num_str @{separate_1_bdv_n),
  1420 		(*"?bdv ^^^ ?n / ?b = 1 / ?b * ?bdv ^^^ ?n"*)
  1421 		Thm ("nadd_divide_distrib", 
  1422 		     num_str @{thm nadd_divide_distrib})
  1423 		(*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"
  1424 		      WN051031 DOES NOT BELONG TO HERE*)
  1425 		];
  1426 val make_ratpoly_in = prep_rls(
  1427   Seq {id = "make_ratpoly_in", preconds = []:term list, 
  1428        rew_ord = ("dummy_ord", dummy_ord),
  1429       erls = Atools_erls, srls = Erls,
  1430       calc = [], (*asm_thm = [],*)
  1431       rules = [Rls_ norm_Rational,
  1432 	       Rls_ order_add_mult_in,
  1433 	       Rls_ discard_parentheses,
  1434 	       Rls_ separate_bdvs,
  1435 	       (* Rls_ rearrange_assoc, WN060916 why does cancel_p not work?*)
  1436 	       Rls_ cancel_p
  1437 	       (*Calc ("HOL.divide"  ,eval_cancel "#divide_") too weak!*)
  1438 	       ],
  1439       scr = EmptyScr}:rls);      
  1440 
  1441 
  1442 ruleset' := overwritelthy @{theory} (!ruleset',
  1443   [("order_add_mult_in", order_add_mult_in),
  1444    ("collect_bdv", collect_bdv),
  1445    ("make_polynomial_in", make_polynomial_in),
  1446    ("make_ratpoly_in", make_ratpoly_in),
  1447    ("separate_bdvs", separate_bdvs)
  1448    ]);
  1449 *}
  1450 
  1451 end
  1452 
  1453 
  1454 
  1455 
  1456 
  1457