src/Tools/isac/Knowledge/PolyEq.thy
author wneuper <walther.neuper@jku.at>
Tue, 10 Aug 2021 09:43:07 +0200
changeset 60358 8377b6c37640
parent 60342 e96abd81a321
child 60360 49680d595342
permissions -rw-r--r--
complete replacement of Rule.Thm/Eval by \<^rule_thm> and \<^rule_eval> in src/*
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
walther@60358
   304
  bdv_n_collect_assoc1_1: "l * bdv \<up> n + (m * bdv \<up> n + k) = (l + m) * bdv \<up> n + k" and
walther@60242
   305
  bdv_n_collect_assoc1_2: "bdv \<up> n + (m * bdv \<up> n + k) = (1 + m) * bdv \<up> n + k" and
walther@60242
   306
  bdv_n_collect_assoc1_3: "l * bdv \<up> n + (bdv \<up> n + k) = (l + 1) * bdv \<up> n + k" and
neuper@37906
   307
walther@60242
   308
  bdv_n_collect_assoc2_1: "k + l * bdv \<up> n + m * bdv \<up> n = k +(l + m) * bdv \<up> n" and
walther@60242
   309
  bdv_n_collect_assoc2_2: "k + bdv \<up> n + m * bdv \<up> n = k + (1 + m) * bdv \<up> n" and
walther@60242
   310
  bdv_n_collect_assoc2_3: "k + l * bdv \<up> n + bdv \<up> n = k + (l + 1) * bdv \<up> n" and
neuper@37906
   311
neuper@37906
   312
(*WN.14.3.03*)
neuper@52148
   313
  real_minus_div:         "- (a / b) = (-1 * a) / b" and
neuper@38030
   314
                          
neuper@52148
   315
  separate_bdv:           "(a * bdv) / b = (a / b) * (bdv::real)" and
walther@60242
   316
  separate_bdv_n:         "(a * bdv \<up> n) / b = (a / b) * bdv \<up> n" and
neuper@52148
   317
  separate_1_bdv:         "bdv / b = (1 / b) * (bdv::real)" and
walther@60242
   318
  separate_1_bdv_n:       "bdv \<up> n / b = (1 / b) * bdv \<up> n"
neuper@37906
   319
wneuper@59472
   320
ML \<open>
neuper@37972
   321
val thy = @{theory};
neuper@37972
   322
neuper@37954
   323
val PolyEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
walther@60358
   324
  Rule_Set.append_rules "PolyEq_prls" Rule_Set.empty [
walther@60358
   325
     \<^rule_eval>\<open>ident\<close> (Prog_Expr.eval_ident "#ident_"),
walther@60358
   326
     \<^rule_eval>\<open>matches\<close> (Prog_Expr.eval_matches "#matches_"),
walther@60358
   327
     \<^rule_eval>\<open>lhs\<close> (Prog_Expr.eval_lhs ""),
walther@60358
   328
     \<^rule_eval>\<open>rhs\<close> (Prog_Expr.eval_rhs ""),
walther@60358
   329
     \<^rule_eval>\<open>is_expanded_in\<close> (eval_is_expanded_in ""),
walther@60358
   330
     \<^rule_eval>\<open>is_poly_in\<close> (eval_is_poly_in ""),
walther@60358
   331
     \<^rule_eval>\<open>has_degree_in\<close> (eval_has_degree_in ""),    
walther@60358
   332
     \<^rule_eval>\<open>is_polyrat_in\<close> (eval_is_polyrat_in ""),
walther@60358
   333
     \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
walther@60358
   334
     \<^rule_eval>\<open>is_rootTerm_in\<close> (eval_is_rootTerm_in ""),
walther@60358
   335
     \<^rule_eval>\<open>is_ratequation_in\<close> (eval_is_ratequation_in ""),
walther@60358
   336
     \<^rule_thm>\<open>not_true\<close>,
walther@60358
   337
     \<^rule_thm>\<open>not_false\<close>,
walther@60358
   338
     \<^rule_thm>\<open>and_true\<close>,
walther@60358
   339
     \<^rule_thm>\<open>and_false\<close>,
walther@60358
   340
     \<^rule_thm>\<open>or_true\<close>,
walther@60358
   341
     \<^rule_thm>\<open>or_false\<close>];
neuper@37954
   342
neuper@37954
   343
val PolyEq_erls = 
walther@60358
   344
  Rule_Set.merge "PolyEq_erls" LinEq_erls
walther@60358
   345
    (Rule_Set.append_rules "ops_preds" calculate_Rational [
walther@60358
   346
     \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
wenzelm@60297
   347
		 \<^rule_thm>\<open>plus_leq\<close>,
wenzelm@60297
   348
		 \<^rule_thm>\<open>minus_leq\<close>,
wenzelm@60297
   349
		 \<^rule_thm>\<open>rat_leq1\<close>,
wenzelm@60297
   350
		 \<^rule_thm>\<open>rat_leq2\<close>,
walther@60358
   351
		 \<^rule_thm>\<open>rat_leq3\<close>]);
neuper@37954
   352
neuper@37954
   353
val PolyEq_crls = 
walther@59852
   354
    Rule_Set.merge "PolyEq_crls" LinEq_crls
walther@60358
   355
    (Rule_Set.append_rules "ops_preds" calculate_Rational [
walther@60358
   356
     \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
wenzelm@60297
   357
		 \<^rule_thm>\<open>plus_leq\<close>,
wenzelm@60297
   358
		 \<^rule_thm>\<open>minus_leq\<close>,
wenzelm@60297
   359
		 \<^rule_thm>\<open>rat_leq1\<close>,
wenzelm@60297
   360
		 \<^rule_thm>\<open>rat_leq2\<close>,
wenzelm@60297
   361
		 \<^rule_thm>\<open>rat_leq3\<close>
neuper@37954
   362
		 ]);
neuper@37954
   363
s1210629013@55444
   364
val cancel_leading_coeff = prep_rls'(
walther@59851
   365
  Rule_Def.Repeat {id = "cancel_leading_coeff", preconds = [], 
walther@60358
   366
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord),
walther@60358
   367
    erls = PolyEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   368
    rules = [
walther@60358
   369
      \<^rule_thm>\<open>cancel_leading_coeff1\<close>,
walther@60358
   370
      \<^rule_thm>\<open>cancel_leading_coeff2\<close>,
walther@60358
   371
      \<^rule_thm>\<open>cancel_leading_coeff3\<close>,
walther@60358
   372
      \<^rule_thm>\<open>cancel_leading_coeff4\<close>,
walther@60358
   373
      \<^rule_thm>\<open>cancel_leading_coeff5\<close>,
walther@60358
   374
      \<^rule_thm>\<open>cancel_leading_coeff6\<close>,
walther@60358
   375
      \<^rule_thm>\<open>cancel_leading_coeff7\<close>,
walther@60358
   376
      \<^rule_thm>\<open>cancel_leading_coeff8\<close>,
walther@60358
   377
      \<^rule_thm>\<open>cancel_leading_coeff9\<close>,
walther@60358
   378
      \<^rule_thm>\<open>cancel_leading_coeff10\<close>,
walther@60358
   379
      \<^rule_thm>\<open>cancel_leading_coeff11\<close>,
walther@60358
   380
      \<^rule_thm>\<open>cancel_leading_coeff12\<close>,
walther@60358
   381
      \<^rule_thm>\<open>cancel_leading_coeff13\<close> ],
walther@60358
   382
    scr = Rule.Empty_Prog});
s1210629013@55444
   383
walther@59618
   384
val prep_rls' = Auto_Prog.prep_rls @{theory};
wneuper@59472
   385
\<close>
wneuper@59472
   386
ML\<open>
s1210629013@55444
   387
val complete_square = prep_rls'(
walther@59851
   388
  Rule_Def.Repeat {id = "complete_square", preconds = [], 
walther@60358
   389
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord),
walther@60358
   390
    erls = PolyEq_erls, srls = Rule_Set.Empty, calc = [],  errpatts = [],
walther@60358
   391
    rules = [
walther@60358
   392
      \<^rule_thm>\<open>complete_square1\<close>,
walther@60358
   393
      \<^rule_thm>\<open>complete_square2\<close>,
walther@60358
   394
      \<^rule_thm>\<open>complete_square3\<close>,
walther@60358
   395
      \<^rule_thm>\<open>complete_square4\<close>,
walther@60358
   396
      \<^rule_thm>\<open>complete_square5\<close>],
walther@60358
   397
    scr = Rule.Empty_Prog});
neuper@37954
   398
s1210629013@55444
   399
val polyeq_simplify = prep_rls'(
walther@59851
   400
  Rule_Def.Repeat {id = "polyeq_simplify", preconds = [], 
walther@60358
   401
    rew_ord = ("termlessI",termlessI), 
walther@60358
   402
    erls = PolyEq_erls, 
walther@60358
   403
    srls = Rule_Set.Empty, 
walther@60358
   404
    calc = [], errpatts = [],
walther@60358
   405
    rules = [
walther@60358
   406
      \<^rule_thm>\<open>real_assoc_1\<close>,
walther@60358
   407
  		\<^rule_thm>\<open>real_assoc_2\<close>,
walther@60358
   408
  		\<^rule_thm>\<open>real_diff_minus\<close>,
walther@60358
   409
  		\<^rule_thm>\<open>real_unari_minus\<close>,
walther@60358
   410
  		\<^rule_thm>\<open>realpow_multI\<close>,
walther@60358
   411
  		\<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
walther@60358
   412
  		\<^rule_eval>\<open>minus\<close> (**)(eval_binop "#sub_"),
walther@60358
   413
  		\<^rule_eval>\<open>times\<close> (**)(eval_binop "#mult_"),
walther@60358
   414
  		\<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
walther@60358
   415
  		\<^rule_eval>\<open>sqrt\<close> (eval_sqrt "#sqrt_"),
walther@60358
   416
  		\<^rule_eval>\<open>powr\<close> (**)(eval_binop "#power_"),
walther@60358
   417
       Rule.Rls_ reduce_012],
walther@60358
   418
    scr = Rule.Empty_Prog});
wneuper@59472
   419
\<close>
wenzelm@60289
   420
rule_set_knowledge
wenzelm@60286
   421
  cancel_leading_coeff = cancel_leading_coeff and
wenzelm@60286
   422
  complete_square = complete_square and
wenzelm@60286
   423
  PolyEq_erls = PolyEq_erls and
wenzelm@60286
   424
  polyeq_simplify = polyeq_simplify
wneuper@59472
   425
ML\<open>
neuper@37954
   426
walther@60358
   427
(* the subsequent rule-sets are caused by the lack of rewriting at the time of implementation *)
neuper@37954
   428
(* -- d0 -- *)
neuper@37954
   429
(*isolate the bound variable in an d0 equation; 'bdv' is a meta-constant*)
s1210629013@55444
   430
val d0_polyeq_simplify = prep_rls'(
walther@59851
   431
  Rule_Def.Repeat {id = "d0_polyeq_simplify", preconds = [],
walther@60358
   432
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord),
walther@60358
   433
    erls = PolyEq_erls,
walther@60358
   434
    srls = Rule_Set.Empty, 
walther@60358
   435
    calc = [], errpatts = [],
walther@60358
   436
    rules = [
walther@60358
   437
      \<^rule_thm>\<open>d0_true\<close>,
walther@60358
   438
      \<^rule_thm>\<open>d0_false\<close>],
walther@60358
   439
    scr = Rule.Empty_Prog});
neuper@37954
   440
neuper@37954
   441
(* -- d1 -- *)
neuper@37954
   442
(*isolate the bound variable in an d1 equation; 'bdv' is a meta-constant*)
s1210629013@55444
   443
val d1_polyeq_simplify = prep_rls'(
walther@59851
   444
  Rule_Def.Repeat {id = "d1_polyeq_simplify", preconds = [],
walther@60358
   445
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord),
walther@60358
   446
    erls = PolyEq_erls,
walther@60358
   447
    srls = Rule_Set.Empty, 
walther@60358
   448
    calc = [], errpatts = [],
walther@60358
   449
    rules = [
walther@60358
   450
  		\<^rule_thm>\<open>d1_isolate_add1\<close>, (* a+bx=0 -> bx=-a *)
walther@60358
   451
  		\<^rule_thm>\<open>d1_isolate_add2\<close>, (* a+ x=0 ->  x=-a *)
walther@60358
   452
  		\<^rule_thm>\<open>d1_isolate_div\<close>   (*   bx=c -> x=c/b *)],
walther@60358
   453
    scr = Rule.Empty_Prog});
walther@60358
   454
\<close>
neuper@37954
   455
wneuper@59472
   456
subsection \<open>degree 2\<close>
wneuper@59472
   457
ML\<open>
neuper@42394
   458
(* isolate the bound variable in an d2 equation with bdv only;
neuper@42394
   459
  "bdv" is a meta-constant substituted for the "x" below by isac's rewriter. *)
s1210629013@55444
   460
val d2_polyeq_bdv_only_simplify = prep_rls'(
walther@59857
   461
  Rule_Def.Repeat {id = "d2_polyeq_bdv_only_simplify", preconds = [], rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord),
walther@59851
   462
    erls = PolyEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   463
    rules = [
walther@60358
   464
       \<^rule_thm>\<open>d2_prescind1\<close>, (*   ax+bx^2=0 -> x(a+bx)=0 *)
wenzelm@60297
   465
       \<^rule_thm>\<open>d2_prescind2\<close>, (*   ax+ x^2=0 -> x(a+ x)=0 *)
wenzelm@60297
   466
       \<^rule_thm>\<open>d2_prescind3\<close>, (*    x+bx^2=0 -> x(1+bx)=0 *)
wenzelm@60297
   467
       \<^rule_thm>\<open>d2_prescind4\<close>, (*    x+ x^2=0 -> x(1+ x)=0 *)
wenzelm@60297
   468
       \<^rule_thm>\<open>d2_sqrt_equation1\<close>,    (* x^2=c   -> x=+-sqrt(c) *)
wenzelm@60297
   469
       \<^rule_thm>\<open>d2_sqrt_equation1_neg\<close>, (* [0<c] x^2=c  -> []*)
wenzelm@60297
   470
       \<^rule_thm>\<open>d2_sqrt_equation2\<close>,    (*  x^2=0 ->    x=0       *)
wenzelm@60297
   471
       \<^rule_thm>\<open>d2_reduce_equation1\<close>,(* x(a+bx)=0 -> x=0 |a+bx=0*)
wenzelm@60297
   472
       \<^rule_thm>\<open>d2_reduce_equation2\<close>,(* x(a+ x)=0 -> x=0 |a+ x=0*)
walther@60358
   473
       \<^rule_thm>\<open>d2_isolate_div\<close>],           (* bx^2=c -> x^2=c/b      *)
walther@60358
   474
    scr = Rule.Empty_Prog});
walther@60358
   475
walther@60358
   476
(* isolate the bound variable in an d2 equation with sqrt only;
neuper@37954
   477
   'bdv' is a meta-constant*)
s1210629013@55444
   478
val d2_polyeq_sq_only_simplify = prep_rls'(
walther@59851
   479
  Rule_Def.Repeat {id = "d2_polyeq_sq_only_simplify", preconds = [],
walther@60358
   480
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord),
walther@60358
   481
    erls = PolyEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   482
    rules = [
walther@60358
   483
      \<^rule_thm>\<open>d2_isolate_add1\<close>,(* a+   bx^2=0 -> bx^2=(-1)a*)
walther@60358
   484
  		\<^rule_thm>\<open>d2_isolate_add2\<close>,  (* a+    x^2=0 ->  x^2=(-1)a*)
walther@60358
   485
  		\<^rule_thm>\<open>d2_sqrt_equation2\<close>,     (*  x^2=0 ->    x=0    *)
walther@60358
   486
  		\<^rule_thm>\<open>d2_sqrt_equation1\<close>,      (* x^2=c   -> x=+-sqrt(c)*)
walther@60358
   487
  		\<^rule_thm>\<open>d2_sqrt_equation1_neg\<close>,(* [c<0] x^2=c  -> x=[] *)
walther@60358
   488
  		\<^rule_thm>\<open>d2_isolate_div\<close>],        (* bx^2=c -> x^2=c/b*)
walther@60358
   489
    scr = Rule.Empty_Prog});
wneuper@59472
   490
\<close>
wneuper@59472
   491
ML\<open>
neuper@37954
   492
(* isolate the bound variable in an d2 equation with pqFormula;
neuper@37954
   493
   'bdv' is a meta-constant*)
s1210629013@55444
   494
val d2_polyeq_pqFormula_simplify = prep_rls'(
walther@59851
   495
  Rule_Def.Repeat {id = "d2_polyeq_pqFormula_simplify", preconds = [],
walther@60358
   496
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord), erls = PolyEq_erls,
walther@60358
   497
    srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   498
    rules = [
walther@60358
   499
      \<^rule_thm>\<open>d2_pqformula1\<close>, (* q+px+ x^2=0 *)
walther@60358
   500
  		\<^rule_thm>\<open>d2_pqformula1_neg\<close>, (* q+px+ x^2=0 *)
walther@60358
   501
  		\<^rule_thm>\<open>d2_pqformula2\<close>, (* q+px+1x^2=0 *)
walther@60358
   502
  		\<^rule_thm>\<open>d2_pqformula2_neg\<close>, (* q+px+1x^2=0 *)
walther@60358
   503
  		\<^rule_thm>\<open>d2_pqformula3\<close>, (* q+ x+ x^2=0 *)
walther@60358
   504
  		\<^rule_thm>\<open>d2_pqformula3_neg\<close>, (* q+ x+ x^2=0 *)
walther@60358
   505
  		\<^rule_thm>\<open>d2_pqformula4\<close>, (* q+ x+1x^2=0 *)
walther@60358
   506
  		\<^rule_thm>\<open>d2_pqformula4_neg\<close>, (* q+ x+1x^2=0 *)
walther@60358
   507
  		\<^rule_thm>\<open>d2_pqformula5\<close>, (*   qx+ x^2=0 *)
walther@60358
   508
  		\<^rule_thm>\<open>d2_pqformula6\<close>, (*   qx+1x^2=0 *)
walther@60358
   509
  		\<^rule_thm>\<open>d2_pqformula7\<close>, (*    x+ x^2=0 *)
walther@60358
   510
  		\<^rule_thm>\<open>d2_pqformula8\<close>, (*    x+1x^2=0 *)
walther@60358
   511
  		\<^rule_thm>\<open>d2_pqformula9\<close>, (* q   +1x^2=0 *)
walther@60358
   512
  		\<^rule_thm>\<open>d2_pqformula9_neg\<close>, (* q   +1x^2=0 *)
walther@60358
   513
  		\<^rule_thm>\<open>d2_pqformula10\<close>, (* q   + x^2=0 *)
walther@60358
   514
  		\<^rule_thm>\<open>d2_pqformula10_neg\<close>, (* q   + x^2=0 *)
walther@60358
   515
  		\<^rule_thm>\<open>d2_sqrt_equation2\<close>, (*       x^2=0 *)
walther@60358
   516
  		\<^rule_thm>\<open>d2_sqrt_equation3\<close>], (*      1x^2=0 *)
walther@60358
   517
    scr = Rule.Empty_Prog});
walther@60358
   518
walther@60358
   519
(* isolate the bound variable in an d2 equation with abcFormula;
neuper@37954
   520
   'bdv' is a meta-constant*)
s1210629013@55444
   521
val d2_polyeq_abcFormula_simplify = prep_rls'(
walther@59851
   522
  Rule_Def.Repeat {id = "d2_polyeq_abcFormula_simplify", preconds = [],
walther@60358
   523
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord), erls = PolyEq_erls,
walther@60358
   524
    srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   525
    rules = [
walther@60358
   526
      \<^rule_thm>\<open>d2_abcformula1\<close>, (*c+bx+cx^2=0 *)
walther@60358
   527
  		\<^rule_thm>\<open>d2_abcformula1_neg\<close>, (*c+bx+cx^2=0 *)
walther@60358
   528
  		\<^rule_thm>\<open>d2_abcformula2\<close>, (*c+ x+cx^2=0 *)
walther@60358
   529
  		\<^rule_thm>\<open>d2_abcformula2_neg\<close>, (*c+ x+cx^2=0 *)
walther@60358
   530
  		\<^rule_thm>\<open>d2_abcformula3\<close>, (*c+bx+ x^2=0 *)
walther@60358
   531
  		\<^rule_thm>\<open>d2_abcformula3_neg\<close>, (*c+bx+ x^2=0 *)
walther@60358
   532
  		\<^rule_thm>\<open>d2_abcformula4\<close>, (*c+ x+ x^2=0 *)
walther@60358
   533
  		\<^rule_thm>\<open>d2_abcformula4_neg\<close>, (*c+ x+ x^2=0 *)
walther@60358
   534
  		\<^rule_thm>\<open>d2_abcformula5\<close>, (*c+   cx^2=0 *)
walther@60358
   535
  		\<^rule_thm>\<open>d2_abcformula5_neg\<close>, (*c+   cx^2=0 *)
walther@60358
   536
  		\<^rule_thm>\<open>d2_abcformula6\<close>, (*c+    x^2=0 *)
walther@60358
   537
  		\<^rule_thm>\<open>d2_abcformula6_neg\<close>, (*c+    x^2=0 *)
walther@60358
   538
  		\<^rule_thm>\<open>d2_abcformula7\<close>, (*  bx+ax^2=0 *)
walther@60358
   539
  		\<^rule_thm>\<open>d2_abcformula8\<close>, (*  bx+ x^2=0 *)
walther@60358
   540
  		\<^rule_thm>\<open>d2_abcformula9\<close>, (*   x+ax^2=0 *)
walther@60358
   541
  		\<^rule_thm>\<open>d2_abcformula10\<close>, (*   x+ x^2=0 *)
walther@60358
   542
  		\<^rule_thm>\<open>d2_sqrt_equation2\<close>, (*      x^2=0 *)  
walther@60358
   543
  		\<^rule_thm>\<open>d2_sqrt_equation3\<close>], (*     bx^2=0 *)  
walther@60358
   544
    scr = Rule.Empty_Prog});
neuper@37954
   545
neuper@37954
   546
(* isolate the bound variable in an d2 equation; 
neuper@37954
   547
   'bdv' is a meta-constant*)
s1210629013@55444
   548
val d2_polyeq_simplify = prep_rls'(
walther@59851
   549
  Rule_Def.Repeat {id = "d2_polyeq_simplify", preconds = [],
walther@60358
   550
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord), erls = PolyEq_erls,
walther@60358
   551
    srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   552
    rules = [
walther@60358
   553
      \<^rule_thm>\<open>d2_pqformula1\<close>, (* p+qx+ x^2=0 *)
walther@60358
   554
  		\<^rule_thm>\<open>d2_pqformula1_neg\<close>, (* p+qx+ x^2=0 *)
walther@60358
   555
  		\<^rule_thm>\<open>d2_pqformula2\<close>, (* p+qx+1x^2=0 *)
walther@60358
   556
  		\<^rule_thm>\<open>d2_pqformula2_neg\<close>, (* p+qx+1x^2=0 *)
walther@60358
   557
  		\<^rule_thm>\<open>d2_pqformula3\<close>, (* p+ x+ x^2=0 *)
walther@60358
   558
  		\<^rule_thm>\<open>d2_pqformula3_neg\<close>, (* p+ x+ x^2=0 *)
walther@60358
   559
  		\<^rule_thm>\<open>d2_pqformula4\<close>, (* p+ x+1x^2=0 *)
walther@60358
   560
  		\<^rule_thm>\<open>d2_pqformula4_neg\<close>, (* p+ x+1x^2=0 *)
walther@60358
   561
  		\<^rule_thm>\<open>d2_abcformula1\<close>, (* c+bx+cx^2=0 *)
walther@60358
   562
  		\<^rule_thm>\<open>d2_abcformula1_neg\<close>, (* c+bx+cx^2=0 *)
walther@60358
   563
  		\<^rule_thm>\<open>d2_abcformula2\<close>, (* c+ x+cx^2=0 *)
walther@60358
   564
  		\<^rule_thm>\<open>d2_abcformula2_neg\<close>, (* c+ x+cx^2=0 *)
walther@60358
   565
  		\<^rule_thm>\<open>d2_prescind1\<close>, (*   ax+bx^2=0 -> x(a+bx)=0 *)
walther@60358
   566
  		\<^rule_thm>\<open>d2_prescind2\<close>, (*   ax+ x^2=0 -> x(a+ x)=0 *)
walther@60358
   567
  		\<^rule_thm>\<open>d2_prescind3\<close>, (*    x+bx^2=0 -> x(1+bx)=0 *)
walther@60358
   568
  		\<^rule_thm>\<open>d2_prescind4\<close>, (*    x+ x^2=0 -> x(1+ x)=0 *)
walther@60358
   569
  		\<^rule_thm>\<open>d2_isolate_add1\<close>, (* a+   bx^2=0 -> bx^2=(-1)a*)
walther@60358
   570
  		\<^rule_thm>\<open>d2_isolate_add2\<close>, (* a+    x^2=0 ->  x^2=(-1)a*)
walther@60358
   571
  		\<^rule_thm>\<open>d2_sqrt_equation1\<close>, (* x^2=c   -> x=+-sqrt(c)*)
walther@60358
   572
  		\<^rule_thm>\<open>d2_sqrt_equation1_neg\<close>, (* [c<0] x^2=c   -> x=[]*)
walther@60358
   573
  		\<^rule_thm>\<open>d2_sqrt_equation2\<close>, (*  x^2=0 ->    x=0    *)
walther@60358
   574
  		\<^rule_thm>\<open>d2_reduce_equation1\<close>, (* x(a+bx)=0 -> x=0 | a+bx=0*)
walther@60358
   575
  		\<^rule_thm>\<open>d2_reduce_equation2\<close>, (* x(a+ x)=0 -> x=0 | a+ x=0*)
walther@60358
   576
  		\<^rule_thm>\<open>d2_isolate_div\<close>], (* bx^2=c -> x^2=c/b*)
walther@60358
   577
    scr = Rule.Empty_Prog});
neuper@37954
   578
neuper@37954
   579
(* -- d3 -- *)
neuper@37954
   580
(* isolate the bound variable in an d3 equation; 'bdv' is a meta-constant *)
s1210629013@55444
   581
val d3_polyeq_simplify = prep_rls'(
walther@59851
   582
  Rule_Def.Repeat {id = "d3_polyeq_simplify", preconds = [],
walther@60358
   583
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord), erls = PolyEq_erls,
walther@60358
   584
    srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   585
    rules = [
walther@60358
   586
      \<^rule_thm>\<open>d3_reduce_equation1\<close>, (*a*bdv + b*bdv \<up> 2 + c*bdv \<up> 3=0) = (bdv=0 | (a + b*bdv + c*bdv \<up> 2=0)*)
walther@60358
   587
    	\<^rule_thm>\<open>d3_reduce_equation2\<close>, (*  bdv + b*bdv \<up> 2 + c*bdv \<up> 3=0) = (bdv=0 | (1 + b*bdv + c*bdv \<up> 2=0)*)
walther@60358
   588
    	\<^rule_thm>\<open>d3_reduce_equation3\<close>, (*a*bdv +   bdv \<up> 2 + c*bdv \<up> 3=0) = (bdv=0 | (a +   bdv + c*bdv \<up> 2=0)*)
walther@60358
   589
    	\<^rule_thm>\<open>d3_reduce_equation4\<close>, (*  bdv +   bdv \<up> 2 + c*bdv \<up> 3=0) = (bdv=0 | (1 +   bdv + c*bdv \<up> 2=0)*)
walther@60358
   590
    	\<^rule_thm>\<open>d3_reduce_equation5\<close>, (*a*bdv + b*bdv \<up> 2 +   bdv \<up> 3=0) = (bdv=0 | (a + b*bdv +   bdv \<up> 2=0)*)
walther@60358
   591
    	\<^rule_thm>\<open>d3_reduce_equation6\<close>, (*  bdv + b*bdv \<up> 2 +   bdv \<up> 3=0) = (bdv=0 | (1 + b*bdv +   bdv \<up> 2=0)*)
walther@60358
   592
    	\<^rule_thm>\<open>d3_reduce_equation7\<close>, (*a*bdv +   bdv \<up> 2 +   bdv \<up> 3=0) = (bdv=0 | (1 +   bdv +   bdv \<up> 2=0)*)
walther@60358
   593
    	\<^rule_thm>\<open>d3_reduce_equation8\<close>, (*  bdv +   bdv \<up> 2 +   bdv \<up> 3=0) = (bdv=0 | (1 +   bdv +   bdv \<up> 2=0)*)
walther@60358
   594
    	\<^rule_thm>\<open>d3_reduce_equation9\<close>, (*a*bdv             + c*bdv \<up> 3=0) = (bdv=0 | (a         + c*bdv \<up> 2=0)*)
walther@60358
   595
    	\<^rule_thm>\<open>d3_reduce_equation10\<close>, (*  bdv             + c*bdv \<up> 3=0) = (bdv=0 | (1         + c*bdv \<up> 2=0)*)
walther@60358
   596
    	\<^rule_thm>\<open>d3_reduce_equation11\<close>, (*a*bdv             +   bdv \<up> 3=0) = (bdv=0 | (a         +   bdv \<up> 2=0)*)
walther@60358
   597
    	\<^rule_thm>\<open>d3_reduce_equation12\<close>, (*  bdv             +   bdv \<up> 3=0) = (bdv=0 | (1         +   bdv \<up> 2=0)*)
walther@60358
   598
    	\<^rule_thm>\<open>d3_reduce_equation13\<close>, (*        b*bdv \<up> 2 + c*bdv \<up> 3=0) = (bdv=0 | (    b*bdv + c*bdv \<up> 2=0)*)
walther@60358
   599
    	\<^rule_thm>\<open>d3_reduce_equation14\<close>, (*          bdv \<up> 2 + c*bdv \<up> 3=0) = (bdv=0 | (      bdv + c*bdv \<up> 2=0)*)
walther@60358
   600
    	\<^rule_thm>\<open>d3_reduce_equation15\<close>, (*        b*bdv \<up> 2 +   bdv \<up> 3=0) = (bdv=0 | (    b*bdv +   bdv \<up> 2=0)*)
walther@60358
   601
    	\<^rule_thm>\<open>d3_reduce_equation16\<close>, (*          bdv \<up> 2 +   bdv \<up> 3=0) = (bdv=0 | (      bdv +   bdv \<up> 2=0)*)
walther@60358
   602
    	\<^rule_thm>\<open>d3_isolate_add1\<close>, (*[|Not(bdv occurs_in a)|] ==> (a + b*bdv \<up> 3=0) = (bdv=0 | (b*bdv \<up> 3=a)*)
walther@60358
   603
    	\<^rule_thm>\<open>d3_isolate_add2\<close>, (*[|Not(bdv occurs_in a)|] ==> (a +   bdv \<up> 3=0) = (bdv=0 | (  bdv \<up> 3=a)*)
walther@60358
   604
    	\<^rule_thm>\<open>d3_isolate_div\<close>, (*[|Not(b=0)|] ==> (b*bdv \<up> 3=c) = (bdv \<up> 3=c/b*)
walther@60358
   605
      \<^rule_thm>\<open>d3_root_equation2\<close>, (*(bdv \<up> 3=0) = (bdv=0) *)
walther@60358
   606
    	\<^rule_thm>\<open>d3_root_equation1\<close>], (*bdv \<up> 3=c) = (bdv = nroot 3 c*)
walther@60358
   607
    scr = Rule.Empty_Prog});
neuper@37954
   608
neuper@37954
   609
(* -- d4 -- *)
neuper@37954
   610
(*isolate the bound variable in an d4 equation; 'bdv' is a meta-constant*)
s1210629013@55444
   611
val d4_polyeq_simplify = prep_rls'(
walther@59851
   612
  Rule_Def.Repeat {id = "d4_polyeq_simplify", preconds = [],
walther@60358
   613
    rew_ord = ("e_rew_ord",Rewrite_Ord.e_rew_ord), erls = PolyEq_erls,
walther@60358
   614
    srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   615
    rules = [
walther@60358
   616
      \<^rule_thm>\<open>d4_sub_u1\<close> (* ax^4+bx^2+c=0 -> x=+-sqrt(ax^2+bx^+c) *)],
walther@60358
   617
    scr = Rule.Empty_Prog});
wneuper@59472
   618
\<close>
wenzelm@60289
   619
rule_set_knowledge
wenzelm@60286
   620
  d0_polyeq_simplify = d0_polyeq_simplify and
wenzelm@60286
   621
  d1_polyeq_simplify = d1_polyeq_simplify and
wenzelm@60286
   622
  d2_polyeq_simplify = d2_polyeq_simplify and
wenzelm@60286
   623
  d2_polyeq_bdv_only_simplify = d2_polyeq_bdv_only_simplify and
wenzelm@60286
   624
  d2_polyeq_sq_only_simplify = d2_polyeq_sq_only_simplify and
neuper@52125
   625
wenzelm@60286
   626
  d2_polyeq_pqFormula_simplify = d2_polyeq_pqFormula_simplify and
wenzelm@60286
   627
  d2_polyeq_abcFormula_simplify = d2_polyeq_abcFormula_simplify and
wenzelm@60286
   628
  d3_polyeq_simplify = d3_polyeq_simplify and
wenzelm@60286
   629
  d4_polyeq_simplify = d4_polyeq_simplify
walther@60258
   630
wenzelm@60306
   631
problem pbl_equ_univ_poly : "polynomial/univariate/equation" =
wenzelm@60306
   632
  \<open>PolyEq_prls\<close>
wenzelm@60306
   633
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   634
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   635
  Where:
wenzelm@60306
   636
    "~((e_e::bool) is_ratequation_in (v_v::real))"
wenzelm@60306
   637
	  "~((lhs e_e) is_rootTerm_in (v_v::real))"
wenzelm@60306
   638
	  "~((rhs e_e) is_rootTerm_in (v_v::real))"
wenzelm@60306
   639
  Find: "solutions v_v'i'"
wenzelm@60306
   640
wenzelm@60306
   641
(*--- d0 ---*)
wenzelm@60306
   642
problem pbl_equ_univ_poly_deg0 : "degree_0/polynomial/univariate/equation" =
wenzelm@60306
   643
  \<open>PolyEq_prls\<close>
wenzelm@60306
   644
  Method: "PolyEq/solve_d0_polyeq_equation"
wenzelm@60306
   645
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   646
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   647
  Where:
wenzelm@60306
   648
    "matches (?a = 0) e_e"
wenzelm@60306
   649
    "(lhs e_e) is_poly_in v_v"
wenzelm@60306
   650
    "((lhs e_e) has_degree_in v_v ) = 0"
wenzelm@60306
   651
  Find: "solutions v_v'i'"
wenzelm@60306
   652
wenzelm@60306
   653
(*--- d1 ---*)
wenzelm@60306
   654
problem pbl_equ_univ_poly_deg1 : "degree_1/polynomial/univariate/equation" =
wenzelm@60306
   655
  \<open>PolyEq_prls\<close>
wenzelm@60306
   656
  Method: "PolyEq/solve_d1_polyeq_equation"
wenzelm@60306
   657
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   658
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   659
  Where:
wenzelm@60306
   660
    "matches (?a = 0) e_e"
wenzelm@60306
   661
	  "(lhs e_e) is_poly_in v_v"
wenzelm@60306
   662
	  "((lhs e_e) has_degree_in v_v ) = 1"
wenzelm@60306
   663
  Find: "solutions v_v'i'"
wenzelm@60306
   664
wenzelm@60306
   665
(*--- d2 ---*)
wenzelm@60306
   666
problem pbl_equ_univ_poly_deg2 : "degree_2/polynomial/univariate/equation" =
wenzelm@60306
   667
  \<open>PolyEq_prls\<close>
wenzelm@60306
   668
  Method: "PolyEq/solve_d2_polyeq_equation"
wenzelm@60306
   669
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   670
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   671
  Where:
wenzelm@60306
   672
    "matches (?a = 0) e_e"
wenzelm@60306
   673
    "(lhs e_e) is_poly_in v_v "
wenzelm@60306
   674
    "((lhs e_e) has_degree_in v_v ) = 2"
wenzelm@60306
   675
  Find: "solutions v_v'i'"
wenzelm@60306
   676
wenzelm@60306
   677
problem pbl_equ_univ_poly_deg2_sqonly : "sq_only/degree_2/polynomial/univariate/equation" =
wenzelm@60306
   678
  \<open>PolyEq_prls\<close>
wenzelm@60306
   679
  Method: "PolyEq/solve_d2_polyeq_sqonly_equation"
wenzelm@60306
   680
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   681
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   682
  Where:
wenzelm@60306
   683
    "matches ( ?a +    ?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   684
     matches ( ?a + ?b*?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   685
     matches (         ?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   686
     matches (      ?b*?v_ \<up> 2 = 0) e_e"
wenzelm@60306
   687
    "Not (matches (?a +    ?v_ +    ?v_ \<up> 2 = 0) e_e) &
wenzelm@60306
   688
     Not (matches (?a + ?b*?v_ +    ?v_ \<up> 2 = 0) e_e) &
wenzelm@60306
   689
     Not (matches (?a +    ?v_ + ?c*?v_ \<up> 2 = 0) e_e) &
wenzelm@60306
   690
     Not (matches (?a + ?b*?v_ + ?c*?v_ \<up> 2 = 0) e_e) &
wenzelm@60306
   691
     Not (matches (        ?v_ +    ?v_ \<up> 2 = 0) e_e) &
wenzelm@60306
   692
     Not (matches (     ?b*?v_ +    ?v_ \<up> 2 = 0) e_e) &
wenzelm@60306
   693
     Not (matches (        ?v_ + ?c*?v_ \<up> 2 = 0) e_e) &
wenzelm@60306
   694
     Not (matches (     ?b*?v_ + ?c*?v_ \<up> 2 = 0) e_e)"
wenzelm@60306
   695
  Find: "solutions v_v'i'"
wenzelm@60306
   696
wenzelm@60306
   697
problem pbl_equ_univ_poly_deg2_bdvonly : "bdv_only/degree_2/polynomial/univariate/equation" =
wenzelm@60306
   698
  \<open>PolyEq_prls\<close>
wenzelm@60306
   699
  Method: "PolyEq/solve_d2_polyeq_bdvonly_equation"
wenzelm@60306
   700
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   701
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   702
  Where:
wenzelm@60306
   703
    "matches (?a*?v_ +    ?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   704
     matches (   ?v_ +    ?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   705
     matches (   ?v_ + ?b*?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   706
     matches (?a*?v_ + ?b*?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   707
     matches (            ?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   708
     matches (         ?b*?v_ \<up> 2 = 0) e_e "
wenzelm@60306
   709
  Find: "solutions v_v'i'"
wenzelm@60306
   710
wenzelm@60306
   711
problem pbl_equ_univ_poly_deg2_pq : "pqFormula/degree_2/polynomial/univariate/equation" =
wenzelm@60306
   712
  \<open>PolyEq_prls\<close>
wenzelm@60306
   713
  Method: "PolyEq/solve_d2_polyeq_pq_equation"
wenzelm@60306
   714
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   715
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   716
  Where:
wenzelm@60306
   717
    "matches (?a + 1*?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   718
     matches (?a +   ?v_ \<up> 2 = 0) e_e"
wenzelm@60306
   719
  Find: "solutions v_v'i'"
wenzelm@60306
   720
wenzelm@60306
   721
problem pbl_equ_univ_poly_deg2_abc : "abcFormula/degree_2/polynomial/univariate/equation" =
wenzelm@60306
   722
  \<open>PolyEq_prls\<close>
wenzelm@60306
   723
  Method: "PolyEq/solve_d2_polyeq_abc_equation"
wenzelm@60306
   724
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   725
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   726
  Where:
wenzelm@60306
   727
    "matches (?a +    ?v_ \<up> 2 = 0) e_e |
wenzelm@60306
   728
     matches (?a + ?b*?v_ \<up> 2 = 0) e_e"
wenzelm@60306
   729
  Find: "solutions v_v'i'"
wenzelm@60306
   730
wenzelm@60306
   731
(*--- d3 ---*)
wenzelm@60306
   732
problem pbl_equ_univ_poly_deg3 : "degree_3/polynomial/univariate/equation" =
wenzelm@60306
   733
  \<open>PolyEq_prls\<close>
wenzelm@60306
   734
  Method: "PolyEq/solve_d3_polyeq_equation"
wenzelm@60306
   735
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   736
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   737
  Where:
wenzelm@60306
   738
    "matches (?a = 0) e_e"
wenzelm@60306
   739
    "(lhs e_e) is_poly_in v_v"
wenzelm@60306
   740
    "((lhs e_e) has_degree_in v_v) = 3"
wenzelm@60306
   741
  Find: "solutions v_v'i'"
wenzelm@60306
   742
wenzelm@60306
   743
(*--- d4 ---*)
wenzelm@60306
   744
problem pbl_equ_univ_poly_deg4 : "degree_4/polynomial/univariate/equation" =
wenzelm@60306
   745
  \<open>PolyEq_prls\<close>
wenzelm@60306
   746
  (*Method: "PolyEq/solve_d4_polyeq_equation"*)
wenzelm@60306
   747
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   748
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   749
  Where:
wenzelm@60306
   750
    "matches (?a = 0) e_e"
wenzelm@60306
   751
    "(lhs e_e) is_poly_in v_v"
wenzelm@60306
   752
    "((lhs e_e) has_degree_in v_v) = 4"
wenzelm@60306
   753
  Find: "solutions v_v'i'"
wenzelm@60306
   754
wenzelm@60306
   755
(*--- normalise ---*)
wenzelm@60306
   756
problem pbl_equ_univ_poly_norm : "normalise/polynomial/univariate/equation" =
wenzelm@60306
   757
  \<open>PolyEq_prls\<close>
wenzelm@60306
   758
  Method: "PolyEq/normalise_poly"
wenzelm@60306
   759
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   760
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   761
  Where:
wenzelm@60306
   762
    "(Not((matches (?a = 0 ) e_e ))) |
wenzelm@60306
   763
     (Not(((lhs e_e) is_poly_in v_v)))"
wenzelm@60306
   764
  Find: "solutions v_v'i'"
wenzelm@60306
   765
wenzelm@60306
   766
(*-------------------------expanded-----------------------*)
wenzelm@60306
   767
problem "pbl_equ_univ_expand" : "expanded/univariate/equation" =
wenzelm@60306
   768
  \<open>PolyEq_prls\<close>
wenzelm@60306
   769
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   770
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   771
  Where:
wenzelm@60306
   772
    "matches (?a = 0) e_e"
wenzelm@60306
   773
    "(lhs e_e) is_expanded_in v_v "
wenzelm@60306
   774
  Find: "solutions v_v'i'"
wenzelm@60306
   775
wenzelm@60306
   776
(*--- d2 ---*)
wenzelm@60306
   777
problem pbl_equ_univ_expand_deg2 : "degree_2/expanded/univariate/equation" =
wenzelm@60306
   778
  \<open>PolyEq_prls\<close>
wenzelm@60306
   779
  Method: "PolyEq/complete_square"
wenzelm@60306
   780
  CAS: "solve (e_e::bool, v_v)"
wenzelm@60306
   781
  Given: "equality e_e" "solveFor v_v"
wenzelm@60306
   782
  Where: "((lhs e_e) has_degree_in v_v) = 2"
wenzelm@60306
   783
  Find: "solutions v_v'i'"
neuper@37954
   784
wneuper@59472
   785
text \<open>"-------------------------methods-----------------------"\<close>
wenzelm@60303
   786
wenzelm@60303
   787
method met_polyeq : "PolyEq" =
wenzelm@60303
   788
  \<open>{rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
wenzelm@60303
   789
    crls=PolyEq_crls, errpats = [], nrls = norm_Rational}\<close>
wneuper@59545
   790
wneuper@59504
   791
partial_function (tailrec) normalize_poly_eq :: "bool \<Rightarrow> real \<Rightarrow> bool"
wneuper@59504
   792
  where
walther@59635
   793
"normalize_poly_eq e_e v_v = (
walther@59635
   794
  let
walther@59635
   795
    e_e = (
walther@59637
   796
      (Try (Rewrite ''all_left'')) #>
walther@59637
   797
      (Try (Repeat (Rewrite ''makex1_x''))) #>
walther@59637
   798
      (Try (Repeat (Rewrite_Set ''expand_binoms''))) #>
walther@59637
   799
      (Try (Repeat (Rewrite_Set_Inst [(''bdv'', v_v)] ''make_ratpoly_in''))) #>
walther@59635
   800
      (Try (Repeat (Rewrite_Set ''polyeq_simplify''))) ) e_e
walther@59635
   801
  in
walther@59635
   802
    SubProblem (''PolyEq'', [''polynomial'', ''univariate'', ''equation''], [''no_met''])
wneuper@59504
   803
      [BOOL e_e, REAL v_v])"
wenzelm@60303
   804
wenzelm@60303
   805
method met_polyeq_norm : "PolyEq/normalise_poly" =
wenzelm@60303
   806
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, calc=[],
wenzelm@60303
   807
    crls=PolyEq_crls, errpats = [], nrls = norm_Rational}\<close>
wenzelm@60303
   808
  Program: normalize_poly_eq.simps
wenzelm@60303
   809
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   810
  Where: "(Not((matches (?a = 0 ) e_e ))) | (Not(((lhs e_e) is_poly_in v_v)))"
wenzelm@60303
   811
  Find: "solutions v_v'i'"
wneuper@59545
   812
wneuper@59504
   813
partial_function (tailrec) solve_poly_equ :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   814
  where
walther@59635
   815
"solve_poly_equ e_e v_v = (
walther@59635
   816
  let
walther@59635
   817
    e_e = (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d0_polyeq_simplify'')) e_e   
walther@59635
   818
  in
walther@59635
   819
    Or_to_List e_e)"
wenzelm@60303
   820
wenzelm@60303
   821
method met_polyeq_d0 : "PolyEq/solve_d0_polyeq_equation" =
wenzelm@60303
   822
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls,
wenzelm@60309
   823
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
   824
    nrls = norm_Rational}\<close>
wenzelm@60303
   825
  Program: solve_poly_equ.simps
wenzelm@60303
   826
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   827
  Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 0"
wenzelm@60303
   828
  Find: "solutions v_v'i'"
wneuper@59545
   829
wneuper@59504
   830
partial_function (tailrec) solve_poly_eq1 :: "bool \<Rightarrow> real \<Rightarrow> bool list"
walther@59635
   831
  where
walther@59635
   832
"solve_poly_eq1 e_e v_v = (
walther@59635
   833
  let
walther@59635
   834
    e_e = (
walther@59637
   835
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d1_polyeq_simplify'')) #>
walther@59637
   836
      (Try (Rewrite_Set ''polyeq_simplify'')) #> 
walther@59635
   837
      (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e;
walther@59635
   838
    L_L = Or_to_List e_e
walther@59635
   839
  in
walther@59635
   840
    Check_elementwise L_L {(v_v::real). Assumptions})"
wenzelm@60303
   841
wenzelm@60303
   842
method met_polyeq_d1 : "PolyEq/solve_d1_polyeq_equation" =
wenzelm@60303
   843
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls,
wenzelm@60309
   844
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
   845
    nrls = norm_Rational}\<close>
wenzelm@60303
   846
  Program: solve_poly_eq1.simps
wenzelm@60303
   847
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   848
  Where: "(lhs e_e) is_poly_in v_v" "((lhs e_e) has_degree_in v_v) = 1"
wenzelm@60303
   849
  Find: "solutions v_v'i'"
wneuper@59545
   850
wneuper@59504
   851
partial_function (tailrec) solve_poly_equ2 :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   852
  where
walther@59635
   853
"solve_poly_equ2 e_e v_v = (
walther@59635
   854
  let
walther@59635
   855
    e_e = (
walther@59637
   856
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_simplify'')) #>
walther@59637
   857
      (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59637
   858
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d1_polyeq_simplify'')) #>
walther@59637
   859
      (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59635
   860
      (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e;
walther@59635
   861
    L_L =  Or_to_List e_e
walther@59635
   862
  in
walther@59635
   863
    Check_elementwise L_L {(v_v::real). Assumptions})"
wenzelm@60303
   864
wenzelm@60303
   865
method met_polyeq_d22 : "PolyEq/solve_d2_polyeq_equation" =
wenzelm@60303
   866
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls,
wenzelm@60309
   867
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
   868
    nrls = norm_Rational}\<close>
wenzelm@60303
   869
  Program: solve_poly_equ2.simps
wenzelm@60303
   870
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   871
  Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2"
wenzelm@60303
   872
  Find: "solutions v_v'i'"
wneuper@59545
   873
wneuper@59504
   874
partial_function (tailrec) solve_poly_equ0 :: "bool \<Rightarrow> real \<Rightarrow> bool list"
walther@59635
   875
  where
walther@59635
   876
"solve_poly_equ0 e_e v_v = (
walther@59635
   877
  let
walther@59635
   878
     e_e = (
walther@59637
   879
       (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_bdv_only_simplify'')) #>
walther@59637
   880
       (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59637
   881
       (Try (Rewrite_Set_Inst [(''bdv'',v_v::real)] ''d1_polyeq_simplify'')) #>
walther@59637
   882
       (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59635
   883
       (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e;
wneuper@59504
   884
     L_L = Or_to_List e_e
walther@59635
   885
  in
walther@59635
   886
    Check_elementwise L_L {(v_v::real). Assumptions})"
wenzelm@60303
   887
wenzelm@60303
   888
method met_polyeq_d2_bdvonly : "PolyEq/solve_d2_polyeq_bdvonly_equation" =
wenzelm@60303
   889
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls,
wenzelm@60309
   890
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
   891
    nrls = norm_Rational}\<close>
wenzelm@60303
   892
  Program: solve_poly_equ0.simps
wenzelm@60303
   893
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   894
  Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2"
wenzelm@60303
   895
  Find: "solutions v_v'i'"
wneuper@59545
   896
wneuper@59504
   897
partial_function (tailrec) solve_poly_equ_sqrt :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   898
  where
walther@59635
   899
"solve_poly_equ_sqrt e_e v_v = (
walther@59635
   900
  let
walther@59635
   901
    e_e = (
walther@59637
   902
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_sq_only_simplify'')) #>
walther@59637
   903
      (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59635
   904
      (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e;
wneuper@59504
   905
    L_L = Or_to_List e_e
walther@59635
   906
  in
walther@59635
   907
    Check_elementwise L_L {(v_v::real). Assumptions})"
wenzelm@60303
   908
wenzelm@60303
   909
method met_polyeq_d2_sqonly : "PolyEq/solve_d2_polyeq_sqonly_equation" =
wenzelm@60303
   910
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls,
wenzelm@60309
   911
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
   912
    nrls = norm_Rational}\<close>
wenzelm@60303
   913
  Program: solve_poly_equ_sqrt.simps
wenzelm@60303
   914
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   915
  Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2"
wenzelm@60303
   916
  Find: "solutions v_v'i'"
wneuper@59545
   917
wneuper@59504
   918
partial_function (tailrec) solve_poly_equ_pq :: "bool \<Rightarrow> real \<Rightarrow> bool list"
walther@59635
   919
  where
walther@59635
   920
"solve_poly_equ_pq e_e v_v = (
walther@59635
   921
  let
walther@59635
   922
    e_e = (
walther@59637
   923
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_pqFormula_simplify'')) #>
walther@59637
   924
      (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59635
   925
      (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e;
walther@59635
   926
    L_L = Or_to_List e_e
walther@59635
   927
  in
walther@59635
   928
    Check_elementwise L_L {(v_v::real). Assumptions})"
wenzelm@60303
   929
wenzelm@60303
   930
method met_polyeq_d2_pq : "PolyEq/solve_d2_polyeq_pq_equation" =
wenzelm@60303
   931
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls,
wenzelm@60309
   932
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
   933
    nrls = norm_Rational}\<close>
wenzelm@60303
   934
  Program: solve_poly_equ_pq.simps
wenzelm@60303
   935
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   936
  Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2"
wenzelm@60303
   937
  Find: "solutions v_v'i'"
wneuper@59545
   938
wneuper@59504
   939
partial_function (tailrec) solve_poly_equ_abc :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   940
  where
walther@59635
   941
"solve_poly_equ_abc e_e v_v = (
walther@59635
   942
  let
walther@59635
   943
    e_e = (
walther@59637
   944
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_abcFormula_simplify'')) #>
walther@59637
   945
      (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59635
   946
      (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e;
walther@59635
   947
    L_L = Or_to_List e_e
wneuper@59504
   948
  in Check_elementwise L_L {(v_v::real). Assumptions})"
wenzelm@60303
   949
wenzelm@60303
   950
method met_polyeq_d2_abc : "PolyEq/solve_d2_polyeq_abc_equation" =
wenzelm@60303
   951
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls,srls=Rule_Set.empty, prls=PolyEq_prls,
wenzelm@60309
   952
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
   953
    nrls = norm_Rational}\<close>
wenzelm@60303
   954
  Program: solve_poly_equ_abc.simps
wenzelm@60303
   955
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   956
  Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2"
wenzelm@60303
   957
  Find: "solutions v_v'i'"
wneuper@59545
   958
wneuper@59504
   959
partial_function (tailrec) solve_poly_equ3 :: "bool \<Rightarrow> real \<Rightarrow> bool list"
walther@59635
   960
  where
walther@59635
   961
"solve_poly_equ3 e_e v_v = (
walther@59635
   962
  let
walther@59635
   963
    e_e = (
walther@59637
   964
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d3_polyeq_simplify'')) #>
walther@59637
   965
      (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59637
   966
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_simplify'')) #>
walther@59637
   967
      (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59637
   968
      (Try (Rewrite_Set_Inst [(''bdv'',v_v)] ''d1_polyeq_simplify'')) #>
walther@59637
   969
      (Try (Rewrite_Set ''polyeq_simplify'')) #>
walther@59635
   970
      (Try (Rewrite_Set ''norm_Rational_parenthesized''))) e_e;
walther@59635
   971
    L_L = Or_to_List e_e
walther@59635
   972
  in
walther@59635
   973
    Check_elementwise L_L {(v_v::real). Assumptions})"
wenzelm@60303
   974
wenzelm@60303
   975
method met_polyeq_d3 : "PolyEq/solve_d3_polyeq_equation" =
wenzelm@60303
   976
  \<open>{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls,
wenzelm@60309
   977
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
   978
    nrls = norm_Rational}\<close>
wenzelm@60303
   979
  Program: solve_poly_equ3.simps
wenzelm@60303
   980
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   981
  Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 3"
wenzelm@60303
   982
  Find: "solutions v_v'i'"
wenzelm@60303
   983
wenzelm@60303
   984
    (*.solves all expanded (ie. normalised) terms of degree 2.*)
s1210629013@55373
   985
    (*Oct.02 restriction: 'eval_true 0 =< discriminant' ony for integer values
s1210629013@55373
   986
      by 'PolyEq_erls'; restricted until Float.thy is implemented*)
wneuper@59504
   987
partial_function (tailrec) solve_by_completing_square :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   988
  where
walther@59635
   989
"solve_by_completing_square e_e v_v = (
walther@59635
   990
  let e_e = (
walther@59637
   991
    (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''cancel_leading_coeff'')) #>
walther@59637
   992
    (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''complete_square'')) #>
walther@59637
   993
    (Try (Rewrite ''square_explicit1'')) #>
walther@59637
   994
    (Try (Rewrite ''square_explicit2'')) #>
walther@59637
   995
    (Rewrite ''root_plus_minus'') #>
walther@59637
   996
    (Try (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''bdv_explicit1''))) #>
walther@59637
   997
    (Try (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''bdv_explicit2''))) #>
walther@59637
   998
    (Try (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''bdv_explicit3''))) #>
walther@59637
   999
    (Try (Rewrite_Set ''calculate_RootRat'')) #>
walther@59635
  1000
    (Try (Repeat (Calculate ''SQRT'')))) e_e
walther@59635
  1001
  in
walther@59635
  1002
    Or_to_List e_e)"
wenzelm@60303
  1003
wenzelm@60303
  1004
method met_polyeq_complsq : "PolyEq/complete_square" =
wenzelm@60303
  1005
  \<open>{rew_ord'="termlessI",rls'=PolyEq_erls,srls=Rule_Set.empty,prls=PolyEq_prls,
wenzelm@60309
  1006
    calc=[("sqrt", (\<^const_name>\<open>sqrt\<close>, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [],
wenzelm@60303
  1007
    nrls = norm_Rational}\<close>
wenzelm@60303
  1008
  Program: solve_by_completing_square.simps
wenzelm@60303
  1009
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
  1010
  Where: "matches (?a = 0) e_e" "((lhs e_e) has_degree_in v_v) = 2"
wenzelm@60303
  1011
  Find: "solutions v_v'i'"
s1210629013@55373
  1012
wneuper@59472
  1013
ML\<open>
neuper@37954
  1014
walther@60342
  1015
(* termorder hacked by MG, adapted later by WN *)
walther@60342
  1016
(**)local (*. for make_polynomial_in .*)
neuper@37954
  1017
neuper@37954
  1018
open Term;  (* for type order = EQUAL | LESS | GREATER *)
neuper@37954
  1019
neuper@37954
  1020
fun pr_ord EQUAL = "EQUAL"
neuper@37954
  1021
  | pr_ord LESS  = "LESS"
neuper@37954
  1022
  | pr_ord GREATER = "GREATER";
neuper@37954
  1023
walther@60263
  1024
fun dest_hd' _ (Const (a, T)) = (((a, 0), T), 0)
walther@60278
  1025
  | dest_hd' x (t as Free (a, T)) =
neuper@37954
  1026
    if x = t then ((("|||||||||||||", 0), T), 0)                        (*WN*)
neuper@37954
  1027
    else (((a, 0), T), 1)
walther@60263
  1028
  | dest_hd' _ (Var v) = (v, 2)
walther@60263
  1029
  | dest_hd' _ (Bound i) = ((("", i), dummyT), 3)
walther@60263
  1030
  | dest_hd' _ (Abs (_, T, _)) = ((("", 0), T), 4)
walther@60263
  1031
  | dest_hd' _ _ = raise ERROR "dest_hd': uncovered case in fun.def.";
neuper@37954
  1032
walther@60342
  1033
fun size_of_term' i pr x (t as Const (\<^const_name>\<open>powr\<close>, _) $ 
walther@60342
  1034
      Free (var, _) $ Free (pot, _)) =
walther@60342
  1035
    (if pr then tracing (idt "#" i ^ "size_of_term' powr: " ^ UnparseC.term t) else ();
walther@60342
  1036
    case x of                                                          (*WN*)
walther@60317
  1037
	    (Free (xstr, _)) => 
walther@60342
  1038
		    if xstr = var 
walther@60342
  1039
        then (if pr then tracing (idt "#" i ^ "xstr = var --> " ^ string_of_int (1000 * (the (TermC.int_opt_of_string pot)))) else ();
walther@60342
  1040
          1000 * (the (TermC.int_opt_of_string pot)))
walther@60342
  1041
        else (if pr then tracing (idt "#" i ^ "x <> Free  --> " ^ "3") else (); 3)
walther@60317
  1042
	  | _ => raise ERROR ("size_of_term' called with subst = " ^ UnparseC.term x))
walther@60342
  1043
  | size_of_term' i pr x (t as Abs (_, _, body)) =
walther@60342
  1044
    (if pr then tracing (idt "#" i ^ "size_of_term' Abs: " ^ UnparseC.term t) else ();
walther@60342
  1045
    1 + size_of_term' (i + 1) pr x body)
walther@60342
  1046
  | size_of_term' i pr x (f $ t) =
walther@60342
  1047
    let
walther@60342
  1048
      val _ = if pr then tracing (idt "#" i ^ "size_of_term' $$$: " ^ UnparseC.term f ^ " $$$ " ^ UnparseC.term t) else ();
walther@60342
  1049
      val s1 = size_of_term' (i + 1) pr x f
walther@60342
  1050
      val s2 = size_of_term' (i + 1) pr x t
walther@60342
  1051
      val _ = if pr then tracing (idt "#" i ^ "size_of_term' $$$-->: " ^ string_of_int s1 ^ " + " ^ string_of_int s2 ^ " = " ^ string_of_int(s1 + s2)) else ();
walther@60342
  1052
    in (s1 + s2) end
walther@60342
  1053
  | size_of_term' i pr x t =
walther@60342
  1054
    (if pr then tracing (idt "#" i ^ "size_of_term' bot: " ^ UnparseC.term t) else ();
walther@60342
  1055
    case t of
walther@60342
  1056
      Free (subst, _) => 
walther@60342
  1057
       (case x of
walther@60342
  1058
   	     Free (xstr, _) =>
walther@60342
  1059
            if xstr = subst
walther@60342
  1060
            then (if pr then tracing (idt "#" i ^ "xstr = var --> " ^ "1000") else (); 1000)
walther@60342
  1061
            else (if pr then tracing (idt "#" i ^ "x <> Free  --> " ^ "1") else (); 1)
walther@60342
  1062
   	   | _ => raise ERROR ("size_of_term' called with subst = " ^ UnparseC.term x))
walther@60342
  1063
     | _ => (if pr then tracing (idt "#" i ^ "bot        --> " ^ "1") else (); 1));
neuper@37954
  1064
walther@60342
  1065
fun term_ord' i pr x thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
walther@60342
  1066
    let
walther@60342
  1067
      val _ = if pr then tracing (idt "#" i ^ "term_ord' Abs") else ();
walther@60342
  1068
      val ord =
walther@60342
  1069
        case term_ord' (i + 1) pr x thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord
walther@60342
  1070
      val _  = if pr then tracing (idt "#" i ^ "term_ord' Abs --> " ^ pr_ord ord) else ()
walther@60342
  1071
    in ord end
walther@60342
  1072
  | term_ord' i pr x _ (t, u) =
walther@60342
  1073
    let
walther@60342
  1074
      val _ = if pr then tracing (idt "#" i ^ "term_ord' bot (" ^ UnparseC.term t ^ ", " ^ UnparseC.term u ^ ")") else ();
walther@60342
  1075
      val ord =
walther@60342
  1076
    	  case int_ord (size_of_term' (i + 1) pr x t, size_of_term' (i + 1) pr x u) of
walther@60342
  1077
    	    EQUAL =>
walther@60342
  1078
    	      let val (f, ts) = strip_comb t and (g, us) = strip_comb u 
walther@60342
  1079
            in
walther@60342
  1080
    	        (case hd_ord (i + 1) pr x (f, g) of 
walther@60342
  1081
    	           EQUAL => (terms_ord x (i + 1) pr) (ts, us) 
walther@60342
  1082
    	         | ord => ord)
walther@60342
  1083
    	      end
walther@60342
  1084
    	  | ord => ord
walther@60342
  1085
      val _  = if pr then tracing (idt "#" i ^ "term_ord' bot --> " ^ pr_ord ord) else ()
walther@60342
  1086
    in ord end
walther@60342
  1087
and hd_ord i pr x (f, g) =                                        (* ~ term.ML *)
walther@60342
  1088
    let
walther@60342
  1089
      val _ = if pr then tracing (idt "#" i ^ "hd_ord (" ^ UnparseC.term f ^ ", " ^ UnparseC.term g ^ ")") else ();
walther@60342
  1090
      val ord = prod_ord
walther@60342
  1091
        (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord
walther@60342
  1092
          (dest_hd' x f, dest_hd' x g)
walther@60342
  1093
      val _ = if pr then tracing (idt "#" i ^ "hd_ord --> " ^ pr_ord ord) else ();
walther@60342
  1094
    in ord end
walther@60342
  1095
and terms_ord x i pr (ts, us) = 
walther@60342
  1096
    let
walther@60342
  1097
      val _ = if pr then tracing (idt "#" i ^ "terms_ord (" ^ UnparseC.terms ts ^ ", " ^ UnparseC.terms us ^ ")") else ();
walther@60342
  1098
      val ord = list_ord (term_ord' (i + 1) pr x (ThyC.get_theory "Isac_Knowledge"))(ts, us);
walther@60342
  1099
      val _ = if pr then tracing (idt "#" i ^ "terms_ord --> " ^ pr_ord ord) else ();
walther@60342
  1100
    in ord end
neuper@52070
  1101
walther@60342
  1102
(**)in(*local*)
neuper@37954
  1103
walther@60324
  1104
fun ord_make_polynomial_in (pr:bool) thy subst (ts, us) =
walther@60342
  1105
  ((** )tracing ("*** subs variable is: " ^ Env.subst2str subst); ( **)
neuper@37954
  1106
	case subst of
walther@60342
  1107
	  (_, x) :: _ =>
walther@60342
  1108
      term_ord' 1 pr x thy (TermC.numerals_to_Free ts, TermC.numerals_to_Free us) = LESS
walther@60263
  1109
	| _ => raise ERROR ("ord_make_polynomial_in called with subst = " ^ Env.subst2str subst))
walther@60263
  1110
neuper@37989
  1111
end;(*local*)
neuper@37954
  1112
wneuper@59472
  1113
\<close>
wneuper@59472
  1114
ML\<open>
s1210629013@55444
  1115
val order_add_mult_in = prep_rls'(
walther@59851
  1116
  Rule_Def.Repeat{id = "order_add_mult_in", preconds = [], 
walther@60358
  1117
    rew_ord = ("ord_make_polynomial_in", ord_make_polynomial_in false @{theory "Poly"}),
walther@60358
  1118
    erls = Rule_Set.empty,srls = Rule_Set.Empty,
walther@60358
  1119
    calc = [], errpatts = [],
walther@60358
  1120
    rules = [
walther@60358
  1121
       \<^rule_thm>\<open>mult.commute\<close>, (* z * w = w * z *)
walther@60358
  1122
	     \<^rule_thm>\<open>real_mult_left_commute\<close>, (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
walther@60358
  1123
	     \<^rule_thm>\<open>mult.assoc\<close>, (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
walther@60358
  1124
	     \<^rule_thm>\<open>add.commute\<close>,	 (*z + w = w + z*)
walther@60358
  1125
	     \<^rule_thm>\<open>add.left_commute\<close>, (*x + (y + z) = y + (x + z)*)
walther@60358
  1126
	     \<^rule_thm>\<open>add.assoc\<close>], (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
walther@60358
  1127
    scr = Rule.Empty_Prog});
neuper@37954
  1128
wneuper@59472
  1129
\<close>
wneuper@59472
  1130
ML\<open>
s1210629013@55444
  1131
val collect_bdv = prep_rls'(
walther@59851
  1132
  Rule_Def.Repeat{id = "collect_bdv", preconds = [], 
walther@60358
  1133
    rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@60358
  1134
    erls = Rule_Set.empty,srls = Rule_Set.Empty,
walther@60358
  1135
    calc = [], errpatts = [],
walther@60358
  1136
    rules = [\<^rule_thm>\<open>bdv_collect_1\<close>,
walther@60358
  1137
	     \<^rule_thm>\<open>bdv_collect_2\<close>,
walther@60358
  1138
	     \<^rule_thm>\<open>bdv_collect_3\<close>,
walther@60358
  1139
    
walther@60358
  1140
	     \<^rule_thm>\<open>bdv_collect_assoc1_1\<close>,
walther@60358
  1141
	     \<^rule_thm>\<open>bdv_collect_assoc1_2\<close>,
walther@60358
  1142
	     \<^rule_thm>\<open>bdv_collect_assoc1_3\<close>,
walther@60358
  1143
    
walther@60358
  1144
	     \<^rule_thm>\<open>bdv_collect_assoc2_1\<close>,
walther@60358
  1145
	     \<^rule_thm>\<open>bdv_collect_assoc2_2\<close>,
walther@60358
  1146
	     \<^rule_thm>\<open>bdv_collect_assoc2_3\<close>,
walther@60358
  1147
    
walther@60358
  1148
    
walther@60358
  1149
	     \<^rule_thm>\<open>bdv_n_collect_1\<close>,
walther@60358
  1150
	     \<^rule_thm>\<open>bdv_n_collect_2\<close>,
walther@60358
  1151
	     \<^rule_thm>\<open>bdv_n_collect_3\<close>,
walther@60358
  1152
    
walther@60358
  1153
	     \<^rule_thm>\<open>bdv_n_collect_assoc1_1\<close>,
walther@60358
  1154
	     \<^rule_thm>\<open>bdv_n_collect_assoc1_2\<close>,
walther@60358
  1155
	     \<^rule_thm>\<open>bdv_n_collect_assoc1_3\<close>,
walther@60358
  1156
    
walther@60358
  1157
	     \<^rule_thm>\<open>bdv_n_collect_assoc2_1\<close>,
walther@60358
  1158
	     \<^rule_thm>\<open>bdv_n_collect_assoc2_2\<close>,
walther@60358
  1159
	     \<^rule_thm>\<open>bdv_n_collect_assoc2_3\<close>],
walther@60358
  1160
    scr = Rule.Empty_Prog});
neuper@37954
  1161
wneuper@59472
  1162
\<close>
wneuper@59472
  1163
ML\<open>
neuper@37954
  1164
(*.transforms an arbitrary term without roots to a polynomial [4] 
neuper@37954
  1165
   according to knowledge/Poly.sml.*) 
s1210629013@55444
  1166
val make_polynomial_in = prep_rls'(
walther@60358
  1167
  Rule_Set.Sequence {
walther@60358
  1168
    id = "make_polynomial_in", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@60358
  1169
    erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
  1170
    rules = [
walther@60358
  1171
      Rule.Rls_ expand_poly,
walther@60358
  1172
	    Rule.Rls_ order_add_mult_in,
walther@60358
  1173
	    Rule.Rls_ simplify_power,
walther@60358
  1174
	    Rule.Rls_ collect_numerals,
walther@60358
  1175
	    Rule.Rls_ reduce_012,
walther@60358
  1176
	    \<^rule_thm>\<open>realpow_oneI\<close>,
walther@60358
  1177
	    Rule.Rls_ discard_parentheses,
walther@60358
  1178
	    Rule.Rls_ collect_bdv],
walther@60358
  1179
    scr = Rule.Empty_Prog});     
neuper@37954
  1180
wneuper@59472
  1181
\<close>
wneuper@59472
  1182
ML\<open>
walther@60358
  1183
val separate_bdvs = Rule_Set.append_rules "separate_bdvs" collect_bdv [
walther@60358
  1184
  \<^rule_thm>\<open>separate_bdv\<close>, (*"?a * ?bdv / ?b = ?a / ?b * ?bdv"*)
walther@60358
  1185
	\<^rule_thm>\<open>separate_bdv_n\<close>,
walther@60358
  1186
	\<^rule_thm>\<open>separate_1_bdv\<close>, (*"?bdv / ?b = (1 / ?b) * ?bdv"*)
walther@60358
  1187
	\<^rule_thm>\<open>separate_1_bdv_n\<close>, (*"?bdv \<up> ?n / ?b = 1 / ?b * ?bdv \<up> ?n"*)
walther@60358
  1188
	\<^rule_thm>\<open>add_divide_distrib\<close> (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"
walther@60358
  1189
	WN051031 DOES NOT BELONG TO HERE*)];
wneuper@59472
  1190
\<close>
wneuper@59472
  1191
ML\<open>
s1210629013@55444
  1192
val make_ratpoly_in = prep_rls'(
walther@60358
  1193
  Rule_Set.Sequence {
walther@60358
  1194
    id = "make_ratpoly_in", preconds = []:term list, rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@60358
  1195
    erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
  1196
    rules = [
walther@60358
  1197
      Rule.Rls_ norm_Rational,
walther@60358
  1198
	    Rule.Rls_ order_add_mult_in,
walther@60358
  1199
	    Rule.Rls_ discard_parentheses,
walther@60358
  1200
	    Rule.Rls_ separate_bdvs,
walther@60358
  1201
	    (* Rule.Rls_ rearrange_assoc, WN060916 why does cancel_p not work?*)
walther@60358
  1202
	    Rule.Rls_ cancel_p
walther@60358
  1203
	    (*\<^rule_eval>\<open>divide\<close> (eval_cancel "#divide_e") too weak!*)],
walther@60358
  1204
    scr = Rule.Empty_Prog});      
wneuper@59472
  1205
\<close>
wenzelm@60289
  1206
rule_set_knowledge
wenzelm@60286
  1207
  order_add_mult_in = order_add_mult_in and
wenzelm@60286
  1208
  collect_bdv = collect_bdv and
wenzelm@60286
  1209
  make_polynomial_in = make_polynomial_in and
wenzelm@60286
  1210
  make_ratpoly_in = make_ratpoly_in and
wenzelm@60286
  1211
  separate_bdvs = separate_bdvs
wenzelm@60286
  1212
ML \<open>
walther@60278
  1213
\<close> ML \<open>
walther@60278
  1214
\<close> ML \<open>
walther@60278
  1215
\<close>
neuper@37906
  1216
end
neuper@37906
  1217
neuper@37906
  1218
neuper@37906
  1219
neuper@37906
  1220
neuper@37906
  1221
neuper@37906
  1222