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