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