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