neuper@37906: (* theory collecting all knowledge neuper@37906: (predicates 'is_rootEq_in', 'is_sqrt_in', 'is_ratEq_in') neuper@37906: for PolynomialEquations. wneuper@59592: alternative dependencies see @{theory "Isac_Knowledge"} neuper@37906: created by: rlang neuper@37906: date: 02.07 neuper@37906: changed by: rlang neuper@37906: last change by: rlang neuper@37906: date: 03.06.03 neuper@37954: (c) by Richard Lang, 2003 neuper@37906: *) neuper@37906: neuper@37954: theory PolyEq imports LinEq RootRatEq begin neuper@37906: neuper@37906: (*-------------------- rules -------------------------------------------------*) walther@60242: (* type real enforced by op " \ " *) neuper@52148: axiomatization where walther@60242: cancel_leading_coeff1: "Not (c =!= 0) ==> (a + b*bdv + c*bdv \ 2 = 0) = walther@60242: (a/c + b/c*bdv + bdv \ 2 = 0)" and walther@60242: cancel_leading_coeff2: "Not (c =!= 0) ==> (a - b*bdv + c*bdv \ 2 = 0) = walther@60242: (a/c - b/c*bdv + bdv \ 2 = 0)" and walther@60242: cancel_leading_coeff3: "Not (c =!= 0) ==> (a + b*bdv - c*bdv \ 2 = 0) = walther@60242: (a/c + b/c*bdv - bdv \ 2 = 0)" and neuper@37906: walther@60242: cancel_leading_coeff4: "Not (c =!= 0) ==> (a + bdv + c*bdv \ 2 = 0) = walther@60242: (a/c + 1/c*bdv + bdv \ 2 = 0)" and walther@60242: cancel_leading_coeff5: "Not (c =!= 0) ==> (a - bdv + c*bdv \ 2 = 0) = walther@60242: (a/c - 1/c*bdv + bdv \ 2 = 0)" and walther@60242: cancel_leading_coeff6: "Not (c =!= 0) ==> (a + bdv - c*bdv \ 2 = 0) = walther@60242: (a/c + 1/c*bdv - bdv \ 2 = 0)" and neuper@37906: walther@60242: cancel_leading_coeff7: "Not (c =!= 0) ==> ( b*bdv + c*bdv \ 2 = 0) = walther@60242: ( b/c*bdv + bdv \ 2 = 0)" and walther@60242: cancel_leading_coeff8: "Not (c =!= 0) ==> ( b*bdv - c*bdv \ 2 = 0) = walther@60242: ( b/c*bdv - bdv \ 2 = 0)" and neuper@37906: walther@60242: cancel_leading_coeff9: "Not (c =!= 0) ==> ( bdv + c*bdv \ 2 = 0) = walther@60242: ( 1/c*bdv + bdv \ 2 = 0)" and walther@60242: cancel_leading_coeff10:"Not (c =!= 0) ==> ( bdv - c*bdv \ 2 = 0) = walther@60242: ( 1/c*bdv - bdv \ 2 = 0)" and neuper@37906: walther@60242: cancel_leading_coeff11:"Not (c =!= 0) ==> (a + b*bdv \ 2 = 0) = walther@60242: (a/b + bdv \ 2 = 0)" and walther@60242: cancel_leading_coeff12:"Not (c =!= 0) ==> (a - b*bdv \ 2 = 0) = walther@60242: (a/b - bdv \ 2 = 0)" and walther@60242: cancel_leading_coeff13:"Not (c =!= 0) ==> ( b*bdv \ 2 = 0) = walther@60242: ( bdv \ 2 = 0/b)" and neuper@37906: walther@60242: complete_square1: "(q + p*bdv + bdv \ 2 = 0) = walther@60242: (q + (p/2 + bdv) \ 2 = (p/2) \ 2)" and walther@60242: complete_square2: "( p*bdv + bdv \ 2 = 0) = walther@60242: ( (p/2 + bdv) \ 2 = (p/2) \ 2)" and walther@60242: complete_square3: "( bdv + bdv \ 2 = 0) = walther@60242: ( (1/2 + bdv) \ 2 = (1/2) \ 2)" and neuper@37906: walther@60242: complete_square4: "(q - p*bdv + bdv \ 2 = 0) = walther@60242: (q + (p/2 - bdv) \ 2 = (p/2) \ 2)" and walther@60242: complete_square5: "(q + p*bdv - bdv \ 2 = 0) = walther@60242: (q + (p/2 - bdv) \ 2 = (p/2) \ 2)" and neuper@37906: walther@60242: square_explicit1: "(a + b \ 2 = c) = ( b \ 2 = c - a)" and walther@60242: square_explicit2: "(a - b \ 2 = c) = (-(b \ 2) = c - a)" and neuper@37906: walther@60242: (*bdv_explicit* required type constrain to real in --- (-8 - 2*x + x \ 2 = 0), by rewriting ---*) neuper@52148: bdv_explicit1: "(a + bdv = b) = (bdv = - a + (b::real))" and neuper@52148: bdv_explicit2: "(a - bdv = b) = ((-1)*bdv = - a + (b::real))" and neuper@52148: bdv_explicit3: "((-1)*bdv = b) = (bdv = (-1)*(b::real))" and neuper@37906: neuper@52148: plus_leq: "(0 <= a + b) = ((-1)*b <= a)"(*Isa?*) and neuper@52148: minus_leq: "(0 <= a - b) = ( b <= a)"(*Isa?*) and neuper@37906: wneuper@59370: (*-- normalise --*) neuper@37906: (*WN0509 compare LinEq.all_left "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)"*) neuper@52148: all_left: "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)" and walther@60242: makex1_x: "a\1 = a" and neuper@52148: real_assoc_1: "a+(b+c) = a+b+c" and neuper@52148: real_assoc_2: "a*(b*c) = a*b*c" and neuper@37906: neuper@37906: (* ---- degree 0 ----*) neuper@52148: d0_true: "(0=0) = True" and neuper@52148: d0_false: "[|Not(bdv occurs_in a);Not(a=0)|] ==> (a=0) = False" and neuper@37906: (* ---- degree 1 ----*) neuper@37983: d1_isolate_add1: neuper@52148: "[|Not(bdv occurs_in a)|] ==> (a + b*bdv = 0) = (b*bdv = (-1)*a)" and neuper@37983: d1_isolate_add2: neuper@52148: "[|Not(bdv occurs_in a)|] ==> (a + bdv = 0) = ( bdv = (-1)*a)" and neuper@37983: d1_isolate_div: neuper@52148: "[|Not(b=0);Not(bdv occurs_in c)|] ==> (b*bdv = c) = (bdv = c/b)" and neuper@37906: (* ---- degree 2 ----*) neuper@37983: d2_isolate_add1: walther@60242: "[|Not(bdv occurs_in a)|] ==> (a + b*bdv \ 2=0) = (b*bdv \ 2= (-1)*a)" and neuper@37983: d2_isolate_add2: walther@60242: "[|Not(bdv occurs_in a)|] ==> (a + bdv \ 2=0) = ( bdv \ 2= (-1)*a)" and neuper@37983: d2_isolate_div: walther@60242: "[|Not(b=0);Not(bdv occurs_in c)|] ==> (b*bdv \ 2=c) = (bdv \ 2=c/b)" and neuper@42394: walther@60242: d2_prescind1: "(a*bdv + b*bdv \ 2 = 0) = (bdv*(a +b*bdv)=0)" and walther@60242: d2_prescind2: "(a*bdv + bdv \ 2 = 0) = (bdv*(a + bdv)=0)" and walther@60242: d2_prescind3: "( bdv + b*bdv \ 2 = 0) = (bdv*(1+b*bdv)=0)" and walther@60242: d2_prescind4: "( bdv + bdv \ 2 = 0) = (bdv*(1+ bdv)=0)" and neuper@37906: (* eliminate degree 2 *) neuper@37906: (* thm for neg arguments in sqroot have postfix _neg *) neuper@37983: d2_sqrt_equation1: "[|(0<=c);Not(bdv occurs_in c)|] ==> walther@60242: (bdv \ 2=c) = ((bdv=sqrt c) | (bdv=(-1)*sqrt c ))" and t@42197: d2_sqrt_equation1_neg: walther@60242: "[|(c<0);Not(bdv occurs_in c)|] ==> (bdv \ 2=c) = False" and walther@60242: d2_sqrt_equation2: "(bdv \ 2=0) = (bdv=0)" and walther@60242: d2_sqrt_equation3: "(b*bdv \ 2=0) = (bdv=0)" neuper@52148: axiomatization where (*AK..if replaced by "and" we get errors: t@42203: exception PTREE "nth _ []" raised t@42203: (line 783 of "/usr/local/isabisac/src/Tools/isac/Interpret/ctree.sml"): t@42203: 'fun nth _ [] = raise PTREE "nth _ []"' t@42203: and t@42203: exception Bind raised t@42203: (line 1097 of "/usr/local/isabisac/test/Tools/isac/Frontend/interface.sml"): t@42203: 'val (Form f, tac, asms) = pt_extract (pt, p);' *) walther@60242: (* WN120315 these 2 thms need "::real", because no " \ " constrains type as neuper@42394: required in test --- rls d2_polyeq_bdv_only_simplify --- *) neuper@52148: d2_reduce_equation1: "(bdv*(a +b*bdv)=0) = ((bdv=0)|(a+b*bdv=(0::real)))" and neuper@42394: d2_reduce_equation2: "(bdv*(a + bdv)=0) = ((bdv=0)|(a+ bdv=(0::real)))" neuper@52148: neuper@52148: axiomatization where (*..if replaced by "and" we get errors: t@42203: exception PTREE "nth _ []" raised t@42203: (line 783 of "/usr/local/isabisac/src/Tools/isac/Interpret/ctree.sml"): t@42203: 'fun nth _ [] = raise PTREE "nth _ []"' t@42203: and t@42203: exception Bind raised t@42203: (line 1097 of "/usr/local/isabisac/test/Tools/isac/Frontend/interface.sml"): t@42203: 'val (Form f, tac, asms) = pt_extract (pt, p);' *) walther@60273: d2_pqformula1: "[|0<=p \ 2 - 4*q|] ==> (q+p*bdv+ bdv \ 2=0) = walther@60242: ((bdv= (-1)*(p/2) + sqrt(p \ 2 - 4*q)/2) walther@60242: | (bdv= (-1)*(p/2) - sqrt(p \ 2 - 4*q)/2))" and walther@60242: d2_pqformula1_neg: "[|p \ 2 - 4*q<0|] ==> (q+p*bdv+ bdv \ 2=0) = False" and walther@60242: d2_pqformula2: "[|0<=p \ 2 - 4*q|] ==> (q+p*bdv+1*bdv \ 2=0) = walther@60242: ((bdv= (-1)*(p/2) + sqrt(p \ 2 - 4*q)/2) walther@60242: | (bdv= (-1)*(p/2) - sqrt(p \ 2 - 4*q)/2))" and walther@60242: d2_pqformula2_neg: "[|p \ 2 - 4*q<0|] ==> (q+p*bdv+1*bdv \ 2=0) = False" and walther@60242: d2_pqformula3: "[|0<=1 - 4*q|] ==> (q+ bdv+ bdv \ 2=0) = neuper@37954: ((bdv= (-1)*(1/2) + sqrt(1 - 4*q)/2) neuper@52148: | (bdv= (-1)*(1/2) - sqrt(1 - 4*q)/2))" and walther@60242: d2_pqformula3_neg: "[|1 - 4*q<0|] ==> (q+ bdv+ bdv \ 2=0) = False" and walther@60242: d2_pqformula4: "[|0<=1 - 4*q|] ==> (q+ bdv+1*bdv \ 2=0) = neuper@37954: ((bdv= (-1)*(1/2) + sqrt(1 - 4*q)/2) neuper@52148: | (bdv= (-1)*(1/2) - sqrt(1 - 4*q)/2))" and walther@60242: d2_pqformula4_neg: "[|1 - 4*q<0|] ==> (q+ bdv+1*bdv \ 2=0) = False" and walther@60242: d2_pqformula5: "[|0<=p \ 2 - 0|] ==> ( p*bdv+ bdv \ 2=0) = walther@60242: ((bdv= (-1)*(p/2) + sqrt(p \ 2 - 0)/2) walther@60242: | (bdv= (-1)*(p/2) - sqrt(p \ 2 - 0)/2))" and t@42203: (* d2_pqformula5_neg not need p^2 never less zero in R *) walther@60402: d2_pqformula6: " ( p*bdv+1*bdv \ 2=0) = walther@60242: ((bdv= (-1)*(p/2) + sqrt(p \ 2 - 0)/2) walther@60242: | (bdv= (-1)*(p/2) - sqrt(p \ 2 - 0)/2))" and neuper@37906: (* d2_pqformula6_neg not need p^2 never less zero in R *) walther@60402: d2_pqformula7: " ( bdv+ bdv \ 2=0) = neuper@37954: ((bdv= (-1)*(1/2) + sqrt(1 - 0)/2) neuper@52148: | (bdv= (-1)*(1/2) - sqrt(1 - 0)/2))" and neuper@37906: (* d2_pqformula7_neg not need, because 1<0 ==> False*) walther@60402: d2_pqformula8: " ( bdv+1*bdv \ 2=0) = neuper@37954: ((bdv= (-1)*(1/2) + sqrt(1 - 0)/2) neuper@52148: | (bdv= (-1)*(1/2) - sqrt(1 - 0)/2))" and neuper@37906: (* d2_pqformula8_neg not need, because 1<0 ==> False*) neuper@37983: d2_pqformula9: "[|Not(bdv occurs_in q); 0<= (-1)*4*q|] ==> walther@60242: (q+ 1*bdv \ 2=0) = ((bdv= 0 + sqrt(0 - 4*q)/2) neuper@52148: | (bdv= 0 - sqrt(0 - 4*q)/2))" and neuper@37983: d2_pqformula9_neg: walther@60242: "[|Not(bdv occurs_in q); (-1)*4*q<0|] ==> (q+ 1*bdv \ 2=0) = False" and neuper@37983: d2_pqformula10: walther@60242: "[|Not(bdv occurs_in q); 0<= (-1)*4*q|] ==> (q+ bdv \ 2=0) = neuper@37906: ((bdv= 0 + sqrt(0 - 4*q)/2) neuper@52148: | (bdv= 0 - sqrt(0 - 4*q)/2))" and neuper@37983: d2_pqformula10_neg: walther@60242: "[|Not(bdv occurs_in q); (-1)*4*q<0|] ==> (q+ bdv \ 2=0) = False" and neuper@37983: d2_abcformula1: walther@60242: "[|0<=b \ 2 - 4*a*c|] ==> (c + b*bdv+a*bdv \ 2=0) = walther@60242: ((bdv=( -b + sqrt(b \ 2 - 4*a*c))/(2*a)) walther@60242: | (bdv=( -b - sqrt(b \ 2 - 4*a*c))/(2*a)))" and neuper@37983: d2_abcformula1_neg: walther@60242: "[|b \ 2 - 4*a*c<0|] ==> (c + b*bdv+a*bdv \ 2=0) = False" and neuper@37983: d2_abcformula2: walther@60242: "[|0<=1 - 4*a*c|] ==> (c+ bdv+a*bdv \ 2=0) = neuper@37906: ((bdv=( -1 + sqrt(1 - 4*a*c))/(2*a)) neuper@52148: | (bdv=( -1 - sqrt(1 - 4*a*c))/(2*a)))" and neuper@37983: d2_abcformula2_neg: walther@60242: "[|1 - 4*a*c<0|] ==> (c+ bdv+a*bdv \ 2=0) = False" and neuper@37983: d2_abcformula3: walther@60242: "[|0<=b \ 2 - 4*1*c|] ==> (c + b*bdv+ bdv \ 2=0) = walther@60242: ((bdv=( -b + sqrt(b \ 2 - 4*1*c))/(2*1)) walther@60242: | (bdv=( -b - sqrt(b \ 2 - 4*1*c))/(2*1)))" and neuper@37983: d2_abcformula3_neg: walther@60242: "[|b \ 2 - 4*1*c<0|] ==> (c + b*bdv+ bdv \ 2=0) = False" and neuper@37983: d2_abcformula4: walther@60242: "[|0<=1 - 4*1*c|] ==> (c + bdv+ bdv \ 2=0) = neuper@37906: ((bdv=( -1 + sqrt(1 - 4*1*c))/(2*1)) neuper@52148: | (bdv=( -1 - sqrt(1 - 4*1*c))/(2*1)))" and neuper@37983: d2_abcformula4_neg: walther@60242: "[|1 - 4*1*c<0|] ==> (c + bdv+ bdv \ 2=0) = False" and neuper@37983: d2_abcformula5: walther@60242: "[|Not(bdv occurs_in c); 0<=0 - 4*a*c|] ==> (c + a*bdv \ 2=0) = neuper@37906: ((bdv=( 0 + sqrt(0 - 4*a*c))/(2*a)) neuper@52148: | (bdv=( 0 - sqrt(0 - 4*a*c))/(2*a)))" and neuper@37983: d2_abcformula5_neg: walther@60242: "[|Not(bdv occurs_in c); 0 - 4*a*c<0|] ==> (c + a*bdv \ 2=0) = False" and neuper@37983: d2_abcformula6: walther@60242: "[|Not(bdv occurs_in c); 0<=0 - 4*1*c|] ==> (c+ bdv \ 2=0) = neuper@37906: ((bdv=( 0 + sqrt(0 - 4*1*c))/(2*1)) neuper@52148: | (bdv=( 0 - sqrt(0 - 4*1*c))/(2*1)))" and neuper@37983: d2_abcformula6_neg: walther@60242: "[|Not(bdv occurs_in c); 0 - 4*1*c<0|] ==> (c+ bdv \ 2=0) = False" and neuper@37983: d2_abcformula7: walther@60413: " ( b*bdv+a*bdv \ 2=0) = walther@60242: ((bdv=( -b + sqrt(b \ 2 - 0))/(2*a)) walther@60242: | (bdv=( -b - sqrt(b \ 2 - 0))/(2*a)))" and neuper@37906: (* d2_abcformula7_neg not need b^2 never less zero in R *) neuper@37983: d2_abcformula8: walther@60413: " ( b*bdv+ bdv \ 2=0) = walther@60242: ((bdv=( -b + sqrt(b \ 2 - 0))/(2*1)) walther@60242: | (bdv=( -b - sqrt(b \ 2 - 0))/(2*1)))" and neuper@37906: (* d2_abcformula8_neg not need b^2 never less zero in R *) neuper@37983: d2_abcformula9: walther@60413: " ( bdv+a*bdv \ 2=0) = neuper@37906: ((bdv=( -1 + sqrt(1 - 0))/(2*a)) neuper@52148: | (bdv=( -1 - sqrt(1 - 0))/(2*a)))" and neuper@37906: (* d2_abcformula9_neg not need, because 1<0 ==> False*) neuper@37983: d2_abcformula10: walther@60413: " ( bdv+ bdv \ 2=0) = neuper@37906: ((bdv=( -1 + sqrt(1 - 0))/(2*1)) neuper@52148: | (bdv=( -1 - sqrt(1 - 0))/(2*1)))" and neuper@37906: (* d2_abcformula10_neg not need, because 1<0 ==> False*) neuper@37906: t@42203: neuper@37906: (* ---- degree 3 ----*) neuper@37983: d3_reduce_equation1: walther@60242: "(a*bdv + b*bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | (a + b*bdv + c*bdv \ 2=0))" and neuper@37983: d3_reduce_equation2: walther@60242: "( bdv + b*bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | (1 + b*bdv + c*bdv \ 2=0))" and neuper@37983: d3_reduce_equation3: walther@60242: "(a*bdv + bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | (a + bdv + c*bdv \ 2=0))" and neuper@37983: d3_reduce_equation4: walther@60242: "( bdv + bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | (1 + bdv + c*bdv \ 2=0))" and neuper@37983: d3_reduce_equation5: walther@60242: "(a*bdv + b*bdv \ 2 + bdv \ 3=0) = (bdv=0 | (a + b*bdv + bdv \ 2=0))" and neuper@37983: d3_reduce_equation6: walther@60242: "( bdv + b*bdv \ 2 + bdv \ 3=0) = (bdv=0 | (1 + b*bdv + bdv \ 2=0))" and neuper@37983: d3_reduce_equation7: walther@60242: "(a*bdv + bdv \ 2 + bdv \ 3=0) = (bdv=0 | (1 + bdv + bdv \ 2=0))" and neuper@37983: d3_reduce_equation8: walther@60242: "( bdv + bdv \ 2 + bdv \ 3=0) = (bdv=0 | (1 + bdv + bdv \ 2=0))" and neuper@37983: d3_reduce_equation9: walther@60242: "(a*bdv + c*bdv \ 3=0) = (bdv=0 | (a + c*bdv \ 2=0))" and neuper@37983: d3_reduce_equation10: walther@60242: "( bdv + c*bdv \ 3=0) = (bdv=0 | (1 + c*bdv \ 2=0))" and neuper@37983: d3_reduce_equation11: walther@60242: "(a*bdv + bdv \ 3=0) = (bdv=0 | (a + bdv \ 2=0))" and neuper@37983: d3_reduce_equation12: walther@60242: "( bdv + bdv \ 3=0) = (bdv=0 | (1 + bdv \ 2=0))" and neuper@37983: d3_reduce_equation13: walther@60242: "( b*bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | ( b*bdv + c*bdv \ 2=0))" and neuper@37983: d3_reduce_equation14: walther@60242: "( bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | ( bdv + c*bdv \ 2=0))" and neuper@37983: d3_reduce_equation15: walther@60242: "( b*bdv \ 2 + bdv \ 3=0) = (bdv=0 | ( b*bdv + bdv \ 2=0))" and neuper@37983: d3_reduce_equation16: walther@60242: "( bdv \ 2 + bdv \ 3=0) = (bdv=0 | ( bdv + bdv \ 2=0))" and neuper@37983: d3_isolate_add1: walther@60242: "[|Not(bdv occurs_in a)|] ==> (a + b*bdv \ 3=0) = (b*bdv \ 3= (-1)*a)" and neuper@37983: d3_isolate_add2: walther@60242: "[|Not(bdv occurs_in a)|] ==> (a + bdv \ 3=0) = ( bdv \ 3= (-1)*a)" and neuper@37983: d3_isolate_div: walther@60242: "[|Not(b=0);Not(bdv occurs_in a)|] ==> (b*bdv \ 3=c) = (bdv \ 3=c/b)" and neuper@37983: d3_root_equation2: walther@60242: "(bdv \ 3=0) = (bdv=0)" and neuper@37983: d3_root_equation1: walther@60242: "(bdv \ 3=c) = (bdv = nroot 3 c)" and neuper@37906: neuper@37906: (* ---- degree 4 ----*) neuper@37906: (* RL03.FIXME es wir nicht getestet ob u>0 *) neuper@37989: d4_sub_u1: walther@60242: "(c+b*bdv \ 2+a*bdv \ 4=0) = walther@60242: ((a*u \ 2+b*u+c=0) & (bdv \ 2=u))" and neuper@37906: neuper@37906: (* ---- 7.3.02 von Termorder ---- *) neuper@37906: neuper@52148: bdv_collect_1: "l * bdv + m * bdv = (l + m) * bdv" and neuper@52148: bdv_collect_2: "bdv + m * bdv = (1 + m) * bdv" and neuper@52148: bdv_collect_3: "l * bdv + bdv = (l + 1) * bdv" and neuper@37906: neuper@37906: (* bdv_collect_assoc0_1 "l * bdv + m * bdv + k = (l + m) * bdv + k" neuper@37906: bdv_collect_assoc0_2 "bdv + m * bdv + k = (1 + m) * bdv + k" neuper@37906: bdv_collect_assoc0_3 "l * bdv + bdv + k = (l + 1) * bdv + k" neuper@37906: *) neuper@52148: bdv_collect_assoc1_1: "l * bdv + (m * bdv + k) = (l + m) * bdv + k" and neuper@52148: bdv_collect_assoc1_2: "bdv + (m * bdv + k) = (1 + m) * bdv + k" and neuper@52148: bdv_collect_assoc1_3: "l * bdv + (bdv + k) = (l + 1) * bdv + k" and neuper@38030: neuper@52148: bdv_collect_assoc2_1: "k + l * bdv + m * bdv = k + (l + m) * bdv" and neuper@52148: bdv_collect_assoc2_2: "k + bdv + m * bdv = k + (1 + m) * bdv" and neuper@52148: bdv_collect_assoc2_3: "k + l * bdv + bdv = k + (l + 1) * bdv" and neuper@37906: neuper@37906: walther@60242: bdv_n_collect_1: "l * bdv \ n + m * bdv \ n = (l + m) * bdv \ n" and walther@60242: bdv_n_collect_2: " bdv \ n + m * bdv \ n = (1 + m) * bdv \ n" and walther@60242: bdv_n_collect_3: "l * bdv \ n + bdv \ n = (l + 1) * bdv \ n" (*order!*) and neuper@37906: walther@60358: bdv_n_collect_assoc1_1: "l * bdv \ n + (m * bdv \ n + k) = (l + m) * bdv \ n + k" and walther@60242: bdv_n_collect_assoc1_2: "bdv \ n + (m * bdv \ n + k) = (1 + m) * bdv \ n + k" and walther@60242: bdv_n_collect_assoc1_3: "l * bdv \ n + (bdv \ n + k) = (l + 1) * bdv \ n + k" and neuper@37906: walther@60242: bdv_n_collect_assoc2_1: "k + l * bdv \ n + m * bdv \ n = k +(l + m) * bdv \ n" and walther@60242: bdv_n_collect_assoc2_2: "k + bdv \ n + m * bdv \ n = k + (1 + m) * bdv \ n" and walther@60242: bdv_n_collect_assoc2_3: "k + l * bdv \ n + bdv \ n = k + (l + 1) * bdv \ n" and neuper@37906: neuper@37906: (*WN.14.3.03*) neuper@52148: real_minus_div: "- (a / b) = (-1 * a) / b" and neuper@38030: neuper@52148: separate_bdv: "(a * bdv) / b = (a / b) * (bdv::real)" and walther@60242: separate_bdv_n: "(a * bdv \ n) / b = (a / b) * bdv \ n" and neuper@52148: separate_1_bdv: "bdv / b = (1 / b) * (bdv::real)" and walther@60242: separate_1_bdv_n: "bdv \ n / b = (1 / b) * bdv \ n" neuper@37906: wneuper@59472: ML \ neuper@37954: val PolyEq_prls = (*3.10.02:just the following order due to subterm evaluation*) walther@60358: Rule_Set.append_rules "PolyEq_prls" Rule_Set.empty [ walther@60358: \<^rule_eval>\ident\ (Prog_Expr.eval_ident "#ident_"), walther@60358: \<^rule_eval>\matches\ (Prog_Expr.eval_matches "#matches_"), walther@60358: \<^rule_eval>\lhs\ (Prog_Expr.eval_lhs ""), walther@60358: \<^rule_eval>\rhs\ (Prog_Expr.eval_rhs ""), walther@60358: \<^rule_eval>\is_expanded_in\ (eval_is_expanded_in ""), walther@60358: \<^rule_eval>\is_poly_in\ (eval_is_poly_in ""), walther@60358: \<^rule_eval>\has_degree_in\ (eval_has_degree_in ""), walther@60358: \<^rule_eval>\is_polyrat_in\ (eval_is_polyrat_in ""), walther@60358: \<^rule_eval>\HOL.eq\ (Prog_Expr.eval_equal "#equal_"), walther@60358: \<^rule_eval>\is_rootTerm_in\ (eval_is_rootTerm_in ""), walther@60358: \<^rule_eval>\is_ratequation_in\ (eval_is_ratequation_in ""), walther@60358: \<^rule_thm>\not_true\, walther@60358: \<^rule_thm>\not_false\, walther@60358: \<^rule_thm>\and_true\, walther@60358: \<^rule_thm>\and_false\, walther@60358: \<^rule_thm>\or_true\, walther@60358: \<^rule_thm>\or_false\]; neuper@37954: neuper@37954: val PolyEq_erls = walther@60358: Rule_Set.merge "PolyEq_erls" LinEq_erls walther@60358: (Rule_Set.append_rules "ops_preds" calculate_Rational [ walther@60358: \<^rule_eval>\HOL.eq\ (Prog_Expr.eval_equal "#equal_"), wenzelm@60297: \<^rule_thm>\plus_leq\, wenzelm@60297: \<^rule_thm>\minus_leq\, wenzelm@60297: \<^rule_thm>\rat_leq1\, wenzelm@60297: \<^rule_thm>\rat_leq2\, walther@60358: \<^rule_thm>\rat_leq3\]); neuper@37954: neuper@37954: val PolyEq_crls = walther@59852: Rule_Set.merge "PolyEq_crls" LinEq_crls walther@60358: (Rule_Set.append_rules "ops_preds" calculate_Rational [ walther@60358: \<^rule_eval>\HOL.eq\ (Prog_Expr.eval_equal "#equal_"), wenzelm@60297: \<^rule_thm>\plus_leq\, wenzelm@60297: \<^rule_thm>\minus_leq\, wenzelm@60297: \<^rule_thm>\rat_leq1\, wenzelm@60297: \<^rule_thm>\rat_leq2\, wenzelm@60297: \<^rule_thm>\rat_leq3\ neuper@37954: ]); neuper@37954: s1210629013@55444: val cancel_leading_coeff = prep_rls'( walther@59851: Rule_Def.Repeat {id = "cancel_leading_coeff", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), walther@60358: erls = PolyEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\cancel_leading_coeff1\, walther@60358: \<^rule_thm>\cancel_leading_coeff2\, walther@60358: \<^rule_thm>\cancel_leading_coeff3\, walther@60358: \<^rule_thm>\cancel_leading_coeff4\, walther@60358: \<^rule_thm>\cancel_leading_coeff5\, walther@60358: \<^rule_thm>\cancel_leading_coeff6\, walther@60358: \<^rule_thm>\cancel_leading_coeff7\, walther@60358: \<^rule_thm>\cancel_leading_coeff8\, walther@60358: \<^rule_thm>\cancel_leading_coeff9\, walther@60358: \<^rule_thm>\cancel_leading_coeff10\, walther@60358: \<^rule_thm>\cancel_leading_coeff11\, walther@60358: \<^rule_thm>\cancel_leading_coeff12\, walther@60358: \<^rule_thm>\cancel_leading_coeff13\ ], walther@60358: scr = Rule.Empty_Prog}); s1210629013@55444: walther@59618: val prep_rls' = Auto_Prog.prep_rls @{theory}; wneuper@59472: \ wneuper@59472: ML\ s1210629013@55444: val complete_square = prep_rls'( walther@59851: Rule_Def.Repeat {id = "complete_square", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), walther@60358: erls = PolyEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\complete_square1\, walther@60358: \<^rule_thm>\complete_square2\, walther@60358: \<^rule_thm>\complete_square3\, walther@60358: \<^rule_thm>\complete_square4\, walther@60358: \<^rule_thm>\complete_square5\], walther@60358: scr = Rule.Empty_Prog}); neuper@37954: s1210629013@55444: val polyeq_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "polyeq_simplify", preconds = [], walther@60358: rew_ord = ("termlessI",termlessI), walther@60358: erls = PolyEq_erls, walther@60358: srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\real_assoc_1\, walther@60358: \<^rule_thm>\real_assoc_2\, walther@60358: \<^rule_thm>\real_diff_minus\, walther@60358: \<^rule_thm>\real_unari_minus\, walther@60358: \<^rule_thm>\realpow_multI\, walther@60358: \<^rule_eval>\plus\ (**)(eval_binop "#add_"), walther@60358: \<^rule_eval>\minus\ (**)(eval_binop "#sub_"), walther@60358: \<^rule_eval>\times\ (**)(eval_binop "#mult_"), walther@60358: \<^rule_eval>\divide\ (Prog_Expr.eval_cancel "#divide_e"), walther@60358: \<^rule_eval>\sqrt\ (eval_sqrt "#sqrt_"), wenzelm@60405: \<^rule_eval>\realpow\ (**)(eval_binop "#power_"), walther@60358: Rule.Rls_ reduce_012], walther@60358: scr = Rule.Empty_Prog}); wneuper@59472: \ wenzelm@60289: rule_set_knowledge wenzelm@60286: cancel_leading_coeff = cancel_leading_coeff and wenzelm@60286: complete_square = complete_square and wenzelm@60286: PolyEq_erls = PolyEq_erls and wenzelm@60286: polyeq_simplify = polyeq_simplify wneuper@59472: ML\ neuper@37954: walther@60358: (* the subsequent rule-sets are caused by the lack of rewriting at the time of implementation *) neuper@37954: (* -- d0 -- *) neuper@37954: (*isolate the bound variable in an d0 equation; 'bdv' is a meta-constant*) s1210629013@55444: val d0_polyeq_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "d0_polyeq_simplify", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), walther@60358: erls = PolyEq_erls, walther@60358: srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d0_true\, walther@60358: \<^rule_thm>\d0_false\], walther@60358: scr = Rule.Empty_Prog}); neuper@37954: neuper@37954: (* -- d1 -- *) neuper@37954: (*isolate the bound variable in an d1 equation; 'bdv' is a meta-constant*) s1210629013@55444: val d1_polyeq_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "d1_polyeq_simplify", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), walther@60358: erls = PolyEq_erls, walther@60358: srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d1_isolate_add1\, (* a+bx=0 -> bx=-a *) walther@60358: \<^rule_thm>\d1_isolate_add2\, (* a+ x=0 -> x=-a *) walther@60358: \<^rule_thm>\d1_isolate_div\ (* bx=c -> x=c/b *)], walther@60358: scr = Rule.Empty_Prog}); walther@60358: \ neuper@37954: wneuper@59472: subsection \degree 2\ wneuper@59472: ML\ neuper@42394: (* isolate the bound variable in an d2 equation with bdv only; neuper@42394: "bdv" is a meta-constant substituted for the "x" below by isac's rewriter. *) s1210629013@55444: val d2_polyeq_bdv_only_simplify = prep_rls'( Walther@60509: Rule_Def.Repeat {id = "d2_polyeq_bdv_only_simplify", preconds = [], rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), walther@59851: erls = PolyEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d2_prescind1\, (* ax+bx^2=0 -> x(a+bx)=0 *) wenzelm@60297: \<^rule_thm>\d2_prescind2\, (* ax+ x^2=0 -> x(a+ x)=0 *) wenzelm@60297: \<^rule_thm>\d2_prescind3\, (* x+bx^2=0 -> x(1+bx)=0 *) wenzelm@60297: \<^rule_thm>\d2_prescind4\, (* x+ x^2=0 -> x(1+ x)=0 *) wenzelm@60297: \<^rule_thm>\d2_sqrt_equation1\, (* x^2=c -> x=+-sqrt(c) *) wenzelm@60297: \<^rule_thm>\d2_sqrt_equation1_neg\, (* [0 []*) wenzelm@60297: \<^rule_thm>\d2_sqrt_equation2\, (* x^2=0 -> x=0 *) wenzelm@60297: \<^rule_thm>\d2_reduce_equation1\,(* x(a+bx)=0 -> x=0 |a+bx=0*) wenzelm@60297: \<^rule_thm>\d2_reduce_equation2\,(* x(a+ x)=0 -> x=0 |a+ x=0*) walther@60358: \<^rule_thm>\d2_isolate_div\], (* bx^2=c -> x^2=c/b *) walther@60358: scr = Rule.Empty_Prog}); walther@60358: walther@60358: (* isolate the bound variable in an d2 equation with sqrt only; neuper@37954: 'bdv' is a meta-constant*) s1210629013@55444: val d2_polyeq_sq_only_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "d2_polyeq_sq_only_simplify", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), walther@60358: erls = PolyEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d2_isolate_add1\,(* a+ bx^2=0 -> bx^2=(-1)a*) walther@60358: \<^rule_thm>\d2_isolate_add2\, (* a+ x^2=0 -> x^2=(-1)a*) walther@60358: \<^rule_thm>\d2_sqrt_equation2\, (* x^2=0 -> x=0 *) walther@60358: \<^rule_thm>\d2_sqrt_equation1\, (* x^2=c -> x=+-sqrt(c)*) walther@60358: \<^rule_thm>\d2_sqrt_equation1_neg\,(* [c<0] x^2=c -> x=[] *) walther@60358: \<^rule_thm>\d2_isolate_div\], (* bx^2=c -> x^2=c/b*) walther@60358: scr = Rule.Empty_Prog}); wneuper@59472: \ wneuper@59472: ML\ neuper@37954: (* isolate the bound variable in an d2 equation with pqFormula; neuper@37954: 'bdv' is a meta-constant*) s1210629013@55444: val d2_polyeq_pqFormula_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "d2_polyeq_pqFormula_simplify", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), erls = PolyEq_erls, walther@60358: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d2_pqformula1\, (* q+px+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula1_neg\, (* q+px+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula2\, (* q+px+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula2_neg\, (* q+px+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula3\, (* q+ x+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula3_neg\, (* q+ x+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula4\, (* q+ x+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula4_neg\, (* q+ x+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula5\, (* qx+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula6\, (* qx+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula7\, (* x+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula8\, (* x+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula9\, (* q +1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula9_neg\, (* q +1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula10\, (* q + x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula10_neg\, (* q + x^2=0 *) walther@60358: \<^rule_thm>\d2_sqrt_equation2\, (* x^2=0 *) walther@60358: \<^rule_thm>\d2_sqrt_equation3\], (* 1x^2=0 *) walther@60358: scr = Rule.Empty_Prog}); walther@60358: walther@60358: (* isolate the bound variable in an d2 equation with abcFormula; neuper@37954: 'bdv' is a meta-constant*) s1210629013@55444: val d2_polyeq_abcFormula_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "d2_polyeq_abcFormula_simplify", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), erls = PolyEq_erls, walther@60358: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d2_abcformula1\, (*c+bx+cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula1_neg\, (*c+bx+cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula2\, (*c+ x+cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula2_neg\, (*c+ x+cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula3\, (*c+bx+ x^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula3_neg\, (*c+bx+ x^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula4\, (*c+ x+ x^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula4_neg\, (*c+ x+ x^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula5\, (*c+ cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula5_neg\, (*c+ cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula6\, (*c+ x^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula6_neg\, (*c+ x^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula7\, (* bx+ax^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula8\, (* bx+ x^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula9\, (* x+ax^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula10\, (* x+ x^2=0 *) walther@60358: \<^rule_thm>\d2_sqrt_equation2\, (* x^2=0 *) walther@60358: \<^rule_thm>\d2_sqrt_equation3\], (* bx^2=0 *) walther@60358: scr = Rule.Empty_Prog}); neuper@37954: neuper@37954: (* isolate the bound variable in an d2 equation; neuper@37954: 'bdv' is a meta-constant*) s1210629013@55444: val d2_polyeq_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "d2_polyeq_simplify", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), erls = PolyEq_erls, walther@60358: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d2_pqformula1\, (* p+qx+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula1_neg\, (* p+qx+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula2\, (* p+qx+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula2_neg\, (* p+qx+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula3\, (* p+ x+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula3_neg\, (* p+ x+ x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula4\, (* p+ x+1x^2=0 *) walther@60358: \<^rule_thm>\d2_pqformula4_neg\, (* p+ x+1x^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula1\, (* c+bx+cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula1_neg\, (* c+bx+cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula2\, (* c+ x+cx^2=0 *) walther@60358: \<^rule_thm>\d2_abcformula2_neg\, (* c+ x+cx^2=0 *) walther@60358: \<^rule_thm>\d2_prescind1\, (* ax+bx^2=0 -> x(a+bx)=0 *) walther@60358: \<^rule_thm>\d2_prescind2\, (* ax+ x^2=0 -> x(a+ x)=0 *) walther@60358: \<^rule_thm>\d2_prescind3\, (* x+bx^2=0 -> x(1+bx)=0 *) walther@60358: \<^rule_thm>\d2_prescind4\, (* x+ x^2=0 -> x(1+ x)=0 *) walther@60358: \<^rule_thm>\d2_isolate_add1\, (* a+ bx^2=0 -> bx^2=(-1)a*) walther@60358: \<^rule_thm>\d2_isolate_add2\, (* a+ x^2=0 -> x^2=(-1)a*) walther@60358: \<^rule_thm>\d2_sqrt_equation1\, (* x^2=c -> x=+-sqrt(c)*) walther@60358: \<^rule_thm>\d2_sqrt_equation1_neg\, (* [c<0] x^2=c -> x=[]*) walther@60358: \<^rule_thm>\d2_sqrt_equation2\, (* x^2=0 -> x=0 *) walther@60358: \<^rule_thm>\d2_reduce_equation1\, (* x(a+bx)=0 -> x=0 | a+bx=0*) walther@60358: \<^rule_thm>\d2_reduce_equation2\, (* x(a+ x)=0 -> x=0 | a+ x=0*) walther@60358: \<^rule_thm>\d2_isolate_div\], (* bx^2=c -> x^2=c/b*) walther@60358: scr = Rule.Empty_Prog}); neuper@37954: neuper@37954: (* -- d3 -- *) neuper@37954: (* isolate the bound variable in an d3 equation; 'bdv' is a meta-constant *) s1210629013@55444: val d3_polyeq_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "d3_polyeq_simplify", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), erls = PolyEq_erls, walther@60358: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d3_reduce_equation1\, (*a*bdv + b*bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | (a + b*bdv + c*bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation2\, (* bdv + b*bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | (1 + b*bdv + c*bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation3\, (*a*bdv + bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | (a + bdv + c*bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation4\, (* bdv + bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | (1 + bdv + c*bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation5\, (*a*bdv + b*bdv \ 2 + bdv \ 3=0) = (bdv=0 | (a + b*bdv + bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation6\, (* bdv + b*bdv \ 2 + bdv \ 3=0) = (bdv=0 | (1 + b*bdv + bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation7\, (*a*bdv + bdv \ 2 + bdv \ 3=0) = (bdv=0 | (1 + bdv + bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation8\, (* bdv + bdv \ 2 + bdv \ 3=0) = (bdv=0 | (1 + bdv + bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation9\, (*a*bdv + c*bdv \ 3=0) = (bdv=0 | (a + c*bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation10\, (* bdv + c*bdv \ 3=0) = (bdv=0 | (1 + c*bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation11\, (*a*bdv + bdv \ 3=0) = (bdv=0 | (a + bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation12\, (* bdv + bdv \ 3=0) = (bdv=0 | (1 + bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation13\, (* b*bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | ( b*bdv + c*bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation14\, (* bdv \ 2 + c*bdv \ 3=0) = (bdv=0 | ( bdv + c*bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation15\, (* b*bdv \ 2 + bdv \ 3=0) = (bdv=0 | ( b*bdv + bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_reduce_equation16\, (* bdv \ 2 + bdv \ 3=0) = (bdv=0 | ( bdv + bdv \ 2=0)*) walther@60358: \<^rule_thm>\d3_isolate_add1\, (*[|Not(bdv occurs_in a)|] ==> (a + b*bdv \ 3=0) = (bdv=0 | (b*bdv \ 3=a)*) walther@60358: \<^rule_thm>\d3_isolate_add2\, (*[|Not(bdv occurs_in a)|] ==> (a + bdv \ 3=0) = (bdv=0 | ( bdv \ 3=a)*) walther@60358: \<^rule_thm>\d3_isolate_div\, (*[|Not(b=0)|] ==> (b*bdv \ 3=c) = (bdv \ 3=c/b*) walther@60358: \<^rule_thm>\d3_root_equation2\, (*(bdv \ 3=0) = (bdv=0) *) walther@60358: \<^rule_thm>\d3_root_equation1\], (*bdv \ 3=c) = (bdv = nroot 3 c*) walther@60358: scr = Rule.Empty_Prog}); neuper@37954: neuper@37954: (* -- d4 -- *) neuper@37954: (*isolate the bound variable in an d4 equation; 'bdv' is a meta-constant*) s1210629013@55444: val d4_polyeq_simplify = prep_rls'( walther@59851: Rule_Def.Repeat {id = "d4_polyeq_simplify", preconds = [], Walther@60509: rew_ord = ("Rewrite_Ord.id_empty",Rewrite_Ord.function_empty), erls = PolyEq_erls, walther@60358: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\d4_sub_u1\ (* ax^4+bx^2+c=0 -> x=+-sqrt(ax^2+bx^+c) *)], walther@60358: scr = Rule.Empty_Prog}); wneuper@59472: \ wenzelm@60289: rule_set_knowledge wenzelm@60286: d0_polyeq_simplify = d0_polyeq_simplify and wenzelm@60286: d1_polyeq_simplify = d1_polyeq_simplify and wenzelm@60286: d2_polyeq_simplify = d2_polyeq_simplify and wenzelm@60286: d2_polyeq_bdv_only_simplify = d2_polyeq_bdv_only_simplify and wenzelm@60286: d2_polyeq_sq_only_simplify = d2_polyeq_sq_only_simplify and neuper@52125: wenzelm@60286: d2_polyeq_pqFormula_simplify = d2_polyeq_pqFormula_simplify and wenzelm@60286: d2_polyeq_abcFormula_simplify = d2_polyeq_abcFormula_simplify and wenzelm@60286: d3_polyeq_simplify = d3_polyeq_simplify and wenzelm@60286: d4_polyeq_simplify = d4_polyeq_simplify walther@60258: wenzelm@60306: problem pbl_equ_univ_poly : "polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "~((e_e::bool) is_ratequation_in (v_v::real))" wenzelm@60306: "~((lhs e_e) is_rootTerm_in (v_v::real))" wenzelm@60306: "~((rhs e_e) is_rootTerm_in (v_v::real))" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: (*--- d0 ---*) wenzelm@60306: problem pbl_equ_univ_poly_deg0 : "degree_0/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/solve_d0_polyeq_equation" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a = 0) e_e" wenzelm@60306: "(lhs e_e) is_poly_in v_v" wenzelm@60306: "((lhs e_e) has_degree_in v_v ) = 0" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: (*--- d1 ---*) wenzelm@60306: problem pbl_equ_univ_poly_deg1 : "degree_1/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/solve_d1_polyeq_equation" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a = 0) e_e" wenzelm@60306: "(lhs e_e) is_poly_in v_v" wenzelm@60306: "((lhs e_e) has_degree_in v_v ) = 1" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: (*--- d2 ---*) wenzelm@60306: problem pbl_equ_univ_poly_deg2 : "degree_2/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/solve_d2_polyeq_equation" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a = 0) e_e" wenzelm@60306: "(lhs e_e) is_poly_in v_v " wenzelm@60306: "((lhs e_e) has_degree_in v_v ) = 2" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: problem pbl_equ_univ_poly_deg2_sqonly : "sq_only/degree_2/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/solve_d2_polyeq_sqonly_equation" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches ( ?a + ?v_ \ 2 = 0) e_e | wenzelm@60306: matches ( ?a + ?b*?v_ \ 2 = 0) e_e | wenzelm@60306: matches ( ?v_ \ 2 = 0) e_e | wenzelm@60306: matches ( ?b*?v_ \ 2 = 0) e_e" wenzelm@60306: "Not (matches (?a + ?v_ + ?v_ \ 2 = 0) e_e) & wenzelm@60306: Not (matches (?a + ?b*?v_ + ?v_ \ 2 = 0) e_e) & wenzelm@60306: Not (matches (?a + ?v_ + ?c*?v_ \ 2 = 0) e_e) & wenzelm@60306: Not (matches (?a + ?b*?v_ + ?c*?v_ \ 2 = 0) e_e) & wenzelm@60306: Not (matches ( ?v_ + ?v_ \ 2 = 0) e_e) & wenzelm@60306: Not (matches ( ?b*?v_ + ?v_ \ 2 = 0) e_e) & wenzelm@60306: Not (matches ( ?v_ + ?c*?v_ \ 2 = 0) e_e) & wenzelm@60306: Not (matches ( ?b*?v_ + ?c*?v_ \ 2 = 0) e_e)" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: problem pbl_equ_univ_poly_deg2_bdvonly : "bdv_only/degree_2/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/solve_d2_polyeq_bdvonly_equation" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a*?v_ + ?v_ \ 2 = 0) e_e | wenzelm@60306: matches ( ?v_ + ?v_ \ 2 = 0) e_e | wenzelm@60306: matches ( ?v_ + ?b*?v_ \ 2 = 0) e_e | wenzelm@60306: matches (?a*?v_ + ?b*?v_ \ 2 = 0) e_e | wenzelm@60306: matches ( ?v_ \ 2 = 0) e_e | wenzelm@60306: matches ( ?b*?v_ \ 2 = 0) e_e " wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: problem pbl_equ_univ_poly_deg2_pq : "pqFormula/degree_2/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/solve_d2_polyeq_pq_equation" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a + 1*?v_ \ 2 = 0) e_e | wenzelm@60306: matches (?a + ?v_ \ 2 = 0) e_e" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: problem pbl_equ_univ_poly_deg2_abc : "abcFormula/degree_2/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/solve_d2_polyeq_abc_equation" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a + ?v_ \ 2 = 0) e_e | wenzelm@60306: matches (?a + ?b*?v_ \ 2 = 0) e_e" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: (*--- d3 ---*) wenzelm@60306: problem pbl_equ_univ_poly_deg3 : "degree_3/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/solve_d3_polyeq_equation" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a = 0) e_e" wenzelm@60306: "(lhs e_e) is_poly_in v_v" wenzelm@60306: "((lhs e_e) has_degree_in v_v) = 3" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: (*--- d4 ---*) wenzelm@60306: problem pbl_equ_univ_poly_deg4 : "degree_4/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ wenzelm@60306: (*Method: "PolyEq/solve_d4_polyeq_equation"*) wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a = 0) e_e" wenzelm@60306: "(lhs e_e) is_poly_in v_v" wenzelm@60306: "((lhs e_e) has_degree_in v_v) = 4" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: (*--- normalise ---*) wenzelm@60306: problem pbl_equ_univ_poly_norm : "normalise/polynomial/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/normalise_poly" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "(Not((matches (?a = 0 ) e_e ))) | wenzelm@60306: (Not(((lhs e_e) is_poly_in v_v)))" wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: (*-------------------------expanded-----------------------*) wenzelm@60306: problem "pbl_equ_univ_expand" : "expanded/univariate/equation" = wenzelm@60306: \PolyEq_prls\ wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: wenzelm@60306: "matches (?a = 0) e_e" wenzelm@60306: "(lhs e_e) is_expanded_in v_v " wenzelm@60306: Find: "solutions v_v'i'" wenzelm@60306: wenzelm@60306: (*--- d2 ---*) wenzelm@60306: problem pbl_equ_univ_expand_deg2 : "degree_2/expanded/univariate/equation" = wenzelm@60306: \PolyEq_prls\ Walther@60449: Method_Ref: "PolyEq/complete_square" wenzelm@60306: CAS: "solve (e_e::bool, v_v)" wenzelm@60306: Given: "equality e_e" "solveFor v_v" wenzelm@60306: Where: "((lhs e_e) has_degree_in v_v) = 2" wenzelm@60306: Find: "solutions v_v'i'" neuper@37954: wneuper@59472: text \"-------------------------methods-----------------------"\ wenzelm@60303: wenzelm@60303: method met_polyeq : "PolyEq" = wenzelm@60303: \{rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty, wenzelm@60303: crls=PolyEq_crls, errpats = [], nrls = norm_Rational}\ wneuper@59545: wneuper@59504: partial_function (tailrec) normalize_poly_eq :: "bool \ real \ bool" wneuper@59504: where walther@59635: "normalize_poly_eq e_e v_v = ( walther@59635: let walther@59635: e_e = ( walther@59637: (Try (Rewrite ''all_left'')) #> walther@59637: (Try (Repeat (Rewrite ''makex1_x''))) #> walther@59637: (Try (Repeat (Rewrite_Set ''expand_binoms''))) #> walther@59637: (Try (Repeat (Rewrite_Set_Inst [(''bdv'', v_v)] ''make_ratpoly_in''))) #> walther@59635: (Try (Repeat (Rewrite_Set ''polyeq_simplify''))) ) e_e walther@59635: in walther@59635: SubProblem (''PolyEq'', [''polynomial'', ''univariate'', ''equation''], [''no_met'']) wneuper@59504: [BOOL e_e, REAL v_v])" wenzelm@60303: wenzelm@60303: method met_polyeq_norm : "PolyEq/normalise_poly" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, calc=[], wenzelm@60303: crls=PolyEq_crls, errpats = [], nrls = norm_Rational}\ wenzelm@60303: Program: normalize_poly_eq.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(Not((matches (?a = 0 ) e_e ))) | (Not(((lhs e_e) is_poly_in v_v)))" wenzelm@60303: Find: "solutions v_v'i'" wneuper@59545: wneuper@59504: partial_function (tailrec) solve_poly_equ :: "bool \ real \ bool list" wneuper@59504: where walther@59635: "solve_poly_equ e_e v_v = ( walther@59635: let walther@59635: e_e = (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d0_polyeq_simplify'')) e_e walther@59635: in walther@59635: Or_to_List e_e)" wenzelm@60303: wenzelm@60303: method met_polyeq_d0 : "PolyEq/solve_d0_polyeq_equation" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_poly_equ.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 0" wenzelm@60303: Find: "solutions v_v'i'" wneuper@59545: wneuper@59504: partial_function (tailrec) solve_poly_eq1 :: "bool \ real \ bool list" walther@59635: where walther@59635: "solve_poly_eq1 e_e v_v = ( walther@59635: let walther@59635: e_e = ( walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d1_polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59635: (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e; walther@59635: L_L = Or_to_List e_e walther@59635: in walther@59635: Check_elementwise L_L {(v_v::real). Assumptions})" wenzelm@60303: wenzelm@60303: method met_polyeq_d1 : "PolyEq/solve_d1_polyeq_equation" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_poly_eq1.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(lhs e_e) is_poly_in v_v" "((lhs e_e) has_degree_in v_v) = 1" wenzelm@60303: Find: "solutions v_v'i'" wneuper@59545: wneuper@59504: partial_function (tailrec) solve_poly_equ2 :: "bool \ real \ bool list" wneuper@59504: where walther@59635: "solve_poly_equ2 e_e v_v = ( walther@59635: let walther@59635: e_e = ( walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d1_polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59635: (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e; walther@59635: L_L = Or_to_List e_e walther@59635: in walther@59635: Check_elementwise L_L {(v_v::real). Assumptions})" wenzelm@60303: wenzelm@60303: method met_polyeq_d22 : "PolyEq/solve_d2_polyeq_equation" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_poly_equ2.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2" wenzelm@60303: Find: "solutions v_v'i'" wneuper@59545: wneuper@59504: partial_function (tailrec) solve_poly_equ0 :: "bool \ real \ bool list" walther@59635: where walther@59635: "solve_poly_equ0 e_e v_v = ( walther@59635: let walther@59635: e_e = ( walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_bdv_only_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set_Inst [(''bdv'',v_v::real)] ''d1_polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59635: (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e; wneuper@59504: L_L = Or_to_List e_e walther@59635: in walther@59635: Check_elementwise L_L {(v_v::real). Assumptions})" wenzelm@60303: wenzelm@60303: method met_polyeq_d2_bdvonly : "PolyEq/solve_d2_polyeq_bdvonly_equation" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_poly_equ0.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2" wenzelm@60303: Find: "solutions v_v'i'" wneuper@59545: wneuper@59504: partial_function (tailrec) solve_poly_equ_sqrt :: "bool \ real \ bool list" wneuper@59504: where walther@59635: "solve_poly_equ_sqrt e_e v_v = ( walther@59635: let walther@59635: e_e = ( walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_sq_only_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59635: (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e; wneuper@59504: L_L = Or_to_List e_e walther@59635: in walther@59635: Check_elementwise L_L {(v_v::real). Assumptions})" wenzelm@60303: wenzelm@60303: method met_polyeq_d2_sqonly : "PolyEq/solve_d2_polyeq_sqonly_equation" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_poly_equ_sqrt.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2" wenzelm@60303: Find: "solutions v_v'i'" wneuper@59545: wneuper@59504: partial_function (tailrec) solve_poly_equ_pq :: "bool \ real \ bool list" walther@59635: where walther@59635: "solve_poly_equ_pq e_e v_v = ( walther@59635: let walther@59635: e_e = ( walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_pqFormula_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59635: (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e; walther@59635: L_L = Or_to_List e_e walther@59635: in walther@59635: Check_elementwise L_L {(v_v::real). Assumptions})" wenzelm@60303: wenzelm@60303: method met_polyeq_d2_pq : "PolyEq/solve_d2_polyeq_pq_equation" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_poly_equ_pq.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2" wenzelm@60303: Find: "solutions v_v'i'" wneuper@59545: wneuper@59504: partial_function (tailrec) solve_poly_equ_abc :: "bool \ real \ bool list" wneuper@59504: where walther@59635: "solve_poly_equ_abc e_e v_v = ( walther@59635: let walther@59635: e_e = ( walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_abcFormula_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59635: (Try (Rewrite_Set ''norm_Rational_parenthesized'')) ) e_e; walther@59635: L_L = Or_to_List e_e wneuper@59504: in Check_elementwise L_L {(v_v::real). Assumptions})" wenzelm@60303: wenzelm@60303: method met_polyeq_d2_abc : "PolyEq/solve_d2_polyeq_abc_equation" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls,srls=Rule_Set.empty, prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_poly_equ_abc.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 2" wenzelm@60303: Find: "solutions v_v'i'" wneuper@59545: wneuper@59504: partial_function (tailrec) solve_poly_equ3 :: "bool \ real \ bool list" walther@59635: where walther@59635: "solve_poly_equ3 e_e v_v = ( walther@59635: let walther@59635: e_e = ( walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d3_polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''d2_polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set_Inst [(''bdv'',v_v)] ''d1_polyeq_simplify'')) #> walther@59637: (Try (Rewrite_Set ''polyeq_simplify'')) #> walther@59635: (Try (Rewrite_Set ''norm_Rational_parenthesized''))) e_e; walther@59635: L_L = Or_to_List e_e walther@59635: in walther@59635: Check_elementwise L_L {(v_v::real). Assumptions})" wenzelm@60303: wenzelm@60303: method met_polyeq_d3 : "PolyEq/solve_d3_polyeq_equation" = wenzelm@60303: \{rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule_Set.empty, prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_poly_equ3.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "(lhs e_e) is_poly_in v_v " "((lhs e_e) has_degree_in v_v) = 3" wenzelm@60303: Find: "solutions v_v'i'" wenzelm@60303: wenzelm@60303: (*.solves all expanded (ie. normalised) terms of degree 2.*) s1210629013@55373: (*Oct.02 restriction: 'eval_true 0 =< discriminant' ony for integer values s1210629013@55373: by 'PolyEq_erls'; restricted until Float.thy is implemented*) wneuper@59504: partial_function (tailrec) solve_by_completing_square :: "bool \ real \ bool list" wneuper@59504: where walther@59635: "solve_by_completing_square e_e v_v = ( walther@59635: let e_e = ( walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''cancel_leading_coeff'')) #> walther@59637: (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''complete_square'')) #> walther@59637: (Try (Rewrite ''square_explicit1'')) #> walther@59637: (Try (Rewrite ''square_explicit2'')) #> walther@59637: (Rewrite ''root_plus_minus'') #> walther@59637: (Try (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''bdv_explicit1''))) #> walther@59637: (Try (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''bdv_explicit2''))) #> walther@59637: (Try (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''bdv_explicit3''))) #> walther@59637: (Try (Rewrite_Set ''calculate_RootRat'')) #> walther@59635: (Try (Repeat (Calculate ''SQRT'')))) e_e walther@59635: in walther@59635: Or_to_List e_e)" wenzelm@60303: wenzelm@60303: method met_polyeq_complsq : "PolyEq/complete_square" = wenzelm@60303: \{rew_ord'="termlessI",rls'=PolyEq_erls,srls=Rule_Set.empty,prls=PolyEq_prls, wenzelm@60309: calc=[("sqrt", (\<^const_name>\sqrt\, eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], wenzelm@60303: nrls = norm_Rational}\ wenzelm@60303: Program: solve_by_completing_square.simps wenzelm@60303: Given: "equality e_e" "solveFor v_v" wenzelm@60303: Where: "matches (?a = 0) e_e" "((lhs e_e) has_degree_in v_v) = 2" wenzelm@60303: Find: "solutions v_v'i'" s1210629013@55373: wneuper@59472: ML\ neuper@37954: walther@60342: (* termorder hacked by MG, adapted later by WN *) walther@60342: (**)local (*. for make_polynomial_in .*) neuper@37954: neuper@37954: open Term; (* for type order = EQUAL | LESS | GREATER *) neuper@37954: neuper@37954: fun pr_ord EQUAL = "EQUAL" neuper@37954: | pr_ord LESS = "LESS" neuper@37954: | pr_ord GREATER = "GREATER"; neuper@37954: walther@60263: fun dest_hd' _ (Const (a, T)) = (((a, 0), T), 0) walther@60278: | dest_hd' x (t as Free (a, T)) = neuper@37954: if x = t then ((("|||||||||||||", 0), T), 0) (*WN*) neuper@37954: else (((a, 0), T), 1) walther@60263: | dest_hd' _ (Var v) = (v, 2) walther@60263: | dest_hd' _ (Bound i) = ((("", i), dummyT), 3) walther@60263: | dest_hd' _ (Abs (_, T, _)) = ((("", 0), T), 4) walther@60263: | dest_hd' _ _ = raise ERROR "dest_hd': uncovered case in fun.def."; neuper@37954: wenzelm@60405: fun size_of_term' i pr x (t as Const (\<^const_name>\realpow\, _) $ walther@60342: Free (var, _) $ Free (pot, _)) = wenzelm@60405: (if pr then tracing (idt "#" i ^ "size_of_term' realpow: " ^ UnparseC.term t) else (); walther@60342: case x of (*WN*) walther@60317: (Free (xstr, _)) => walther@60342: if xstr = var walther@60342: then (if pr then tracing (idt "#" i ^ "xstr = var --> " ^ string_of_int (1000 * (the (TermC.int_opt_of_string pot)))) else (); walther@60342: 1000 * (the (TermC.int_opt_of_string pot))) walther@60342: else (if pr then tracing (idt "#" i ^ "x <> Free --> " ^ "3") else (); 3) walther@60317: | _ => raise ERROR ("size_of_term' called with subst = " ^ UnparseC.term x)) walther@60342: | size_of_term' i pr x (t as Abs (_, _, body)) = walther@60342: (if pr then tracing (idt "#" i ^ "size_of_term' Abs: " ^ UnparseC.term t) else (); walther@60342: 1 + size_of_term' (i + 1) pr x body) walther@60342: | size_of_term' i pr x (f $ t) = walther@60342: let walther@60342: val _ = if pr then tracing (idt "#" i ^ "size_of_term' $$$: " ^ UnparseC.term f ^ " $$$ " ^ UnparseC.term t) else (); walther@60342: val s1 = size_of_term' (i + 1) pr x f walther@60342: val s2 = size_of_term' (i + 1) pr x t walther@60342: 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: in (s1 + s2) end walther@60342: | size_of_term' i pr x t = walther@60342: (if pr then tracing (idt "#" i ^ "size_of_term' bot: " ^ UnparseC.term t) else (); walther@60342: case t of walther@60342: Free (subst, _) => walther@60342: (case x of walther@60342: Free (xstr, _) => walther@60342: if xstr = subst walther@60342: then (if pr then tracing (idt "#" i ^ "xstr = var --> " ^ "1000") else (); 1000) walther@60342: else (if pr then tracing (idt "#" i ^ "x <> Free --> " ^ "1") else (); 1) walther@60342: | _ => raise ERROR ("size_of_term' called with subst = " ^ UnparseC.term x)) walther@60342: | _ => (if pr then tracing (idt "#" i ^ "bot --> " ^ "1") else (); 1)); neuper@37954: walther@60342: fun term_ord' i pr x thy (Abs (_, T, t), Abs(_, U, u)) = (* ~ term.ML *) walther@60342: let walther@60342: val _ = if pr then tracing (idt "#" i ^ "term_ord' Abs") else (); walther@60342: val ord = walther@60342: case term_ord' (i + 1) pr x thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord walther@60342: val _ = if pr then tracing (idt "#" i ^ "term_ord' Abs --> " ^ pr_ord ord) else () walther@60342: in ord end walther@60342: | term_ord' i pr x _ (t, u) = walther@60342: let walther@60342: val _ = if pr then tracing (idt "#" i ^ "term_ord' bot (" ^ UnparseC.term t ^ ", " ^ UnparseC.term u ^ ")") else (); walther@60342: val ord = walther@60342: case int_ord (size_of_term' (i + 1) pr x t, size_of_term' (i + 1) pr x u) of walther@60342: EQUAL => walther@60342: let val (f, ts) = strip_comb t and (g, us) = strip_comb u walther@60342: in walther@60342: (case hd_ord (i + 1) pr x (f, g) of walther@60342: EQUAL => (terms_ord x (i + 1) pr) (ts, us) walther@60342: | ord => ord) walther@60342: end walther@60342: | ord => ord walther@60342: val _ = if pr then tracing (idt "#" i ^ "term_ord' bot --> " ^ pr_ord ord) else () walther@60342: in ord end walther@60342: and hd_ord i pr x (f, g) = (* ~ term.ML *) walther@60342: let walther@60342: val _ = if pr then tracing (idt "#" i ^ "hd_ord (" ^ UnparseC.term f ^ ", " ^ UnparseC.term g ^ ")") else (); walther@60342: val ord = prod_ord walther@60342: (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord walther@60342: (dest_hd' x f, dest_hd' x g) walther@60342: val _ = if pr then tracing (idt "#" i ^ "hd_ord --> " ^ pr_ord ord) else (); walther@60342: in ord end walther@60342: and terms_ord x i pr (ts, us) = walther@60342: let walther@60342: val _ = if pr then tracing (idt "#" i ^ "terms_ord (" ^ UnparseC.terms ts ^ ", " ^ UnparseC.terms us ^ ")") else (); walther@60342: val ord = list_ord (term_ord' (i + 1) pr x (ThyC.get_theory "Isac_Knowledge"))(ts, us); walther@60342: val _ = if pr then tracing (idt "#" i ^ "terms_ord --> " ^ pr_ord ord) else (); walther@60342: in ord end neuper@52070: walther@60342: (**)in(*local*) neuper@37954: walther@60324: fun ord_make_polynomial_in (pr:bool) thy subst (ts, us) = walther@60342: ((** )tracing ("*** subs variable is: " ^ Env.subst2str subst); ( **) neuper@37954: case subst of walther@60342: (_, x) :: _ => walther@60342: term_ord' 1 pr x thy (TermC.numerals_to_Free ts, TermC.numerals_to_Free us) = LESS walther@60263: | _ => raise ERROR ("ord_make_polynomial_in called with subst = " ^ Env.subst2str subst)) walther@60263: neuper@37989: end;(*local*) neuper@37954: wneuper@59472: \ wneuper@59472: ML\ s1210629013@55444: val order_add_mult_in = prep_rls'( walther@59851: Rule_Def.Repeat{id = "order_add_mult_in", preconds = [], walther@60358: rew_ord = ("ord_make_polynomial_in", ord_make_polynomial_in false @{theory "Poly"}), walther@60358: erls = Rule_Set.empty,srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\mult.commute\, (* z * w = w * z *) walther@60358: \<^rule_thm>\real_mult_left_commute\, (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*) walther@60358: \<^rule_thm>\mult.assoc\, (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*) walther@60358: \<^rule_thm>\add.commute\, (*z + w = w + z*) walther@60358: \<^rule_thm>\add.left_commute\, (*x + (y + z) = y + (x + z)*) walther@60358: \<^rule_thm>\add.assoc\], (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*) walther@60358: scr = Rule.Empty_Prog}); neuper@37954: wneuper@59472: \ wneuper@59472: ML\ s1210629013@55444: val collect_bdv = prep_rls'( walther@59851: Rule_Def.Repeat{id = "collect_bdv", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Rule_Set.empty,srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], walther@60358: rules = [\<^rule_thm>\bdv_collect_1\, walther@60358: \<^rule_thm>\bdv_collect_2\, walther@60358: \<^rule_thm>\bdv_collect_3\, walther@60358: walther@60358: \<^rule_thm>\bdv_collect_assoc1_1\, walther@60358: \<^rule_thm>\bdv_collect_assoc1_2\, walther@60358: \<^rule_thm>\bdv_collect_assoc1_3\, walther@60358: walther@60358: \<^rule_thm>\bdv_collect_assoc2_1\, walther@60358: \<^rule_thm>\bdv_collect_assoc2_2\, walther@60358: \<^rule_thm>\bdv_collect_assoc2_3\, walther@60358: walther@60358: walther@60358: \<^rule_thm>\bdv_n_collect_1\, walther@60358: \<^rule_thm>\bdv_n_collect_2\, walther@60358: \<^rule_thm>\bdv_n_collect_3\, walther@60358: walther@60358: \<^rule_thm>\bdv_n_collect_assoc1_1\, walther@60358: \<^rule_thm>\bdv_n_collect_assoc1_2\, walther@60358: \<^rule_thm>\bdv_n_collect_assoc1_3\, walther@60358: walther@60358: \<^rule_thm>\bdv_n_collect_assoc2_1\, walther@60358: \<^rule_thm>\bdv_n_collect_assoc2_2\, walther@60358: \<^rule_thm>\bdv_n_collect_assoc2_3\], walther@60358: scr = Rule.Empty_Prog}); neuper@37954: wneuper@59472: \ wneuper@59472: ML\ neuper@37954: (*.transforms an arbitrary term without roots to a polynomial [4] neuper@37954: according to knowledge/Poly.sml.*) s1210629013@55444: val make_polynomial_in = prep_rls'( walther@60358: Rule_Set.Sequence { Walther@60509: id = "make_polynomial_in", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: Rule.Rls_ expand_poly, walther@60358: Rule.Rls_ order_add_mult_in, walther@60358: Rule.Rls_ simplify_power, walther@60358: Rule.Rls_ collect_numerals, walther@60358: Rule.Rls_ reduce_012, walther@60358: \<^rule_thm>\realpow_oneI\, walther@60358: Rule.Rls_ discard_parentheses, walther@60358: Rule.Rls_ collect_bdv], walther@60358: scr = Rule.Empty_Prog}); neuper@37954: wneuper@59472: \ wneuper@59472: ML\ walther@60358: val separate_bdvs = Rule_Set.append_rules "separate_bdvs" collect_bdv [ walther@60358: \<^rule_thm>\separate_bdv\, (*"?a * ?bdv / ?b = ?a / ?b * ?bdv"*) walther@60358: \<^rule_thm>\separate_bdv_n\, walther@60358: \<^rule_thm>\separate_1_bdv\, (*"?bdv / ?b = (1 / ?b) * ?bdv"*) walther@60358: \<^rule_thm>\separate_1_bdv_n\, (*"?bdv \ ?n / ?b = 1 / ?b * ?bdv \ ?n"*) walther@60358: \<^rule_thm>\add_divide_distrib\ (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z" walther@60358: WN051031 DOES NOT BELONG TO HERE*)]; wneuper@59472: \ wneuper@59472: ML\ s1210629013@55444: val make_ratpoly_in = prep_rls'( walther@60358: Rule_Set.Sequence { Walther@60509: id = "make_ratpoly_in", preconds = []:term list, rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: Rule.Rls_ norm_Rational, walther@60358: Rule.Rls_ order_add_mult_in, walther@60358: Rule.Rls_ discard_parentheses, walther@60358: Rule.Rls_ separate_bdvs, walther@60358: (* Rule.Rls_ rearrange_assoc, WN060916 why does cancel_p not work?*) walther@60358: Rule.Rls_ cancel_p walther@60358: (*\<^rule_eval>\divide\ (eval_cancel "#divide_e") too weak!*)], walther@60358: scr = Rule.Empty_Prog}); wneuper@59472: \ wenzelm@60289: rule_set_knowledge wenzelm@60286: order_add_mult_in = order_add_mult_in and wenzelm@60286: collect_bdv = collect_bdv and wenzelm@60286: make_polynomial_in = make_polynomial_in and wenzelm@60286: make_ratpoly_in = make_ratpoly_in and wenzelm@60286: separate_bdvs = separate_bdvs wenzelm@60286: ML \ walther@60278: \ ML \ walther@60278: \ ML \ walther@60278: \ neuper@37906: end neuper@37906: neuper@37906: neuper@37906: neuper@37906: neuper@37906: neuper@37906: