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