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