neuper@37906: (* theory collecting all knowledge neuper@37906: (predicates 'is_rootEq_in', 'is_sqrt_in', 'is_ratEq_in') neuper@37906: for PolynomialEquations. neuper@52140: alternative dependencies see @{theory "Isac"} 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: consts neuper@37906: neuper@37906: (*---------scripts--------------------------*) neuper@37906: Complete'_square neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Complete'_square (_ _ =))// (_))" 9) neuper@37906: (*----- poly ----- *) neuper@37906: Normalize'_poly neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Normalize'_poly (_ _=))// (_))" 9) neuper@37906: Solve'_d0'_polyeq'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d0'_polyeq'_equation (_ _ =))// (_))" 9) neuper@37906: Solve'_d1'_polyeq'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d1'_polyeq'_equation (_ _ =))// (_))" 9) neuper@37906: Solve'_d2'_polyeq'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d2'_polyeq'_equation (_ _ =))// (_))" 9) neuper@37906: Solve'_d2'_polyeq'_sqonly'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d2'_polyeq'_sqonly'_equation (_ _ =))// (_))" 9) neuper@37906: Solve'_d2'_polyeq'_bdvonly'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d2'_polyeq'_bdvonly'_equation (_ _ =))// (_))" 9) neuper@37906: Solve'_d2'_polyeq'_pq'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d2'_polyeq'_pq'_equation (_ _ =))// (_))" 9) neuper@37906: Solve'_d2'_polyeq'_abc'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d2'_polyeq'_abc'_equation (_ _ =))// (_))" 9) neuper@37906: Solve'_d3'_polyeq'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d3'_polyeq'_equation (_ _ =))// (_))" 9) neuper@37906: Solve'_d4'_polyeq'_equation neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Solve'_d4'_polyeq'_equation (_ _ =))// (_))" 9) neuper@37906: Biquadrat'_poly neuper@37954: :: "[bool,real, neuper@37954: bool list] => bool list" wneuper@59334: ("((Script Biquadrat'_poly (_ _=))// (_))" 9) neuper@37906: neuper@37906: (*-------------------- rules -------------------------------------------------*) neuper@42394: (* type real enforced by op "^^^" *) neuper@52148: axiomatization where neuper@37983: cancel_leading_coeff1: "Not (c =!= 0) ==> (a + b*bdv + c*bdv^^^2 = 0) = neuper@52148: (a/c + b/c*bdv + bdv^^^2 = 0)" and neuper@37983: cancel_leading_coeff2: "Not (c =!= 0) ==> (a - b*bdv + c*bdv^^^2 = 0) = neuper@52148: (a/c - b/c*bdv + bdv^^^2 = 0)" and neuper@37983: cancel_leading_coeff3: "Not (c =!= 0) ==> (a + b*bdv - c*bdv^^^2 = 0) = neuper@52148: (a/c + b/c*bdv - bdv^^^2 = 0)" and neuper@37906: neuper@37983: cancel_leading_coeff4: "Not (c =!= 0) ==> (a + bdv + c*bdv^^^2 = 0) = neuper@52148: (a/c + 1/c*bdv + bdv^^^2 = 0)" and neuper@37983: cancel_leading_coeff5: "Not (c =!= 0) ==> (a - bdv + c*bdv^^^2 = 0) = neuper@52148: (a/c - 1/c*bdv + bdv^^^2 = 0)" and neuper@37983: cancel_leading_coeff6: "Not (c =!= 0) ==> (a + bdv - c*bdv^^^2 = 0) = neuper@52148: (a/c + 1/c*bdv - bdv^^^2 = 0)" and neuper@37906: neuper@37983: cancel_leading_coeff7: "Not (c =!= 0) ==> ( b*bdv + c*bdv^^^2 = 0) = neuper@52148: ( b/c*bdv + bdv^^^2 = 0)" and neuper@37983: cancel_leading_coeff8: "Not (c =!= 0) ==> ( b*bdv - c*bdv^^^2 = 0) = neuper@52148: ( b/c*bdv - bdv^^^2 = 0)" and neuper@37906: neuper@37983: cancel_leading_coeff9: "Not (c =!= 0) ==> ( bdv + c*bdv^^^2 = 0) = neuper@52148: ( 1/c*bdv + bdv^^^2 = 0)" and neuper@37983: cancel_leading_coeff10:"Not (c =!= 0) ==> ( bdv - c*bdv^^^2 = 0) = neuper@52148: ( 1/c*bdv - bdv^^^2 = 0)" and neuper@37906: neuper@37983: cancel_leading_coeff11:"Not (c =!= 0) ==> (a + b*bdv^^^2 = 0) = neuper@52148: (a/b + bdv^^^2 = 0)" and neuper@37983: cancel_leading_coeff12:"Not (c =!= 0) ==> (a - b*bdv^^^2 = 0) = neuper@52148: (a/b - bdv^^^2 = 0)" and neuper@37983: cancel_leading_coeff13:"Not (c =!= 0) ==> ( b*bdv^^^2 = 0) = neuper@52148: ( bdv^^^2 = 0/b)" and neuper@37906: neuper@37983: complete_square1: "(q + p*bdv + bdv^^^2 = 0) = neuper@52148: (q + (p/2 + bdv)^^^2 = (p/2)^^^2)" and neuper@37983: complete_square2: "( p*bdv + bdv^^^2 = 0) = neuper@52148: ( (p/2 + bdv)^^^2 = (p/2)^^^2)" and neuper@37983: complete_square3: "( bdv + bdv^^^2 = 0) = neuper@52148: ( (1/2 + bdv)^^^2 = (1/2)^^^2)" and neuper@37906: neuper@37983: complete_square4: "(q - p*bdv + bdv^^^2 = 0) = neuper@52148: (q + (p/2 - bdv)^^^2 = (p/2)^^^2)" and neuper@37983: complete_square5: "(q + p*bdv - bdv^^^2 = 0) = neuper@52148: (q + (p/2 - bdv)^^^2 = (p/2)^^^2)" and neuper@37906: neuper@52148: square_explicit1: "(a + b^^^2 = c) = ( b^^^2 = c - a)" and neuper@52148: square_explicit2: "(a - b^^^2 = c) = (-(b^^^2) = c - a)" and neuper@37906: neuper@42318: (*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 neuper@52148: 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: neuper@52148: "[|Not(bdv occurs_in a)|] ==> (a + b*bdv^^^2=0) = (b*bdv^^^2= (-1)*a)" and neuper@37983: d2_isolate_add2: neuper@52148: "[|Not(bdv occurs_in a)|] ==> (a + bdv^^^2=0) = ( bdv^^^2= (-1)*a)" and neuper@37983: d2_isolate_div: neuper@52148: "[|Not(b=0);Not(bdv occurs_in c)|] ==> (b*bdv^^^2=c) = (bdv^^^2=c/b)" and neuper@42394: neuper@52148: d2_prescind1: "(a*bdv + b*bdv^^^2 = 0) = (bdv*(a +b*bdv)=0)" and neuper@52148: d2_prescind2: "(a*bdv + bdv^^^2 = 0) = (bdv*(a + bdv)=0)" and neuper@52148: d2_prescind3: "( bdv + b*bdv^^^2 = 0) = (bdv*(1+b*bdv)=0)" and neuper@52148: 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)|] ==> neuper@52148: (bdv^^^2=c) = ((bdv=sqrt c) | (bdv=(-1)*sqrt c ))" and t@42197: d2_sqrt_equation1_neg: neuper@52148: "[|(c<0);Not(bdv occurs_in c)|] ==> (bdv^^^2=c) = False" and neuper@52148: d2_sqrt_equation2: "(bdv^^^2=0) = (bdv=0)" and neuper@37983: 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);' *) neuper@42394: (* 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);' *) t@42197: d2_pqformula1: "[|0<=p^^^2 - 4*q|] ==> (q+p*bdv+ bdv^^^2=0) = neuper@37954: ((bdv= (-1)*(p/2) + sqrt(p^^^2 - 4*q)/2) neuper@52148: | (bdv= (-1)*(p/2) - sqrt(p^^^2 - 4*q)/2))" and neuper@52148: d2_pqformula1_neg: "[|p^^^2 - 4*q<0|] ==> (q+p*bdv+ bdv^^^2=0) = False" and neuper@37983: d2_pqformula2: "[|0<=p^^^2 - 4*q|] ==> (q+p*bdv+1*bdv^^^2=0) = neuper@37954: ((bdv= (-1)*(p/2) + sqrt(p^^^2 - 4*q)/2) neuper@52148: | (bdv= (-1)*(p/2) - sqrt(p^^^2 - 4*q)/2))" and neuper@52148: d2_pqformula2_neg: "[|p^^^2 - 4*q<0|] ==> (q+p*bdv+1*bdv^^^2=0) = False" and neuper@37983: 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 neuper@52148: d2_pqformula3_neg: "[|1 - 4*q<0|] ==> (q+ bdv+ bdv^^^2=0) = False" and neuper@37983: 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 neuper@52148: d2_pqformula4_neg: "[|1 - 4*q<0|] ==> (q+ bdv+1*bdv^^^2=0) = False" and neuper@37983: d2_pqformula5: "[|0<=p^^^2 - 0|] ==> ( p*bdv+ bdv^^^2=0) = neuper@37954: ((bdv= (-1)*(p/2) + sqrt(p^^^2 - 0)/2) neuper@52148: | (bdv= (-1)*(p/2) - sqrt(p^^^2 - 0)/2))" and t@42203: (* d2_pqformula5_neg not need p^2 never less zero in R *) neuper@37983: d2_pqformula6: "[|0<=p^^^2 - 0|] ==> ( p*bdv+1*bdv^^^2=0) = neuper@37954: ((bdv= (-1)*(p/2) + sqrt(p^^^2 - 0)/2) neuper@52148: | (bdv= (-1)*(p/2) - sqrt(p^^^2 - 0)/2))" and neuper@37906: (* d2_pqformula6_neg not need p^2 never less zero in R *) t@42203: d2_pqformula7: "[|0<=1 - 0|] ==> ( 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*) neuper@37983: d2_pqformula8: "[|0<=1 - 0|] ==> ( 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|] ==> neuper@37954: (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: neuper@52148: "[|Not(bdv occurs_in q); (-1)*4*q<0|] ==> (q+ 1*bdv^^^2=0) = False" and neuper@37983: d2_pqformula10: neuper@37906: "[|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: neuper@52148: "[|Not(bdv occurs_in q); (-1)*4*q<0|] ==> (q+ bdv^^^2=0) = False" and neuper@37983: d2_abcformula1: neuper@37906: "[|0<=b^^^2 - 4*a*c|] ==> (c + b*bdv+a*bdv^^^2=0) = neuper@37906: ((bdv=( -b + sqrt(b^^^2 - 4*a*c))/(2*a)) neuper@52148: | (bdv=( -b - sqrt(b^^^2 - 4*a*c))/(2*a)))" and neuper@37983: d2_abcformula1_neg: neuper@52148: "[|b^^^2 - 4*a*c<0|] ==> (c + b*bdv+a*bdv^^^2=0) = False" and neuper@37983: d2_abcformula2: neuper@37906: "[|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: neuper@52148: "[|1 - 4*a*c<0|] ==> (c+ bdv+a*bdv^^^2=0) = False" and neuper@37983: d2_abcformula3: neuper@37906: "[|0<=b^^^2 - 4*1*c|] ==> (c + b*bdv+ bdv^^^2=0) = neuper@37906: ((bdv=( -b + sqrt(b^^^2 - 4*1*c))/(2*1)) neuper@52148: | (bdv=( -b - sqrt(b^^^2 - 4*1*c))/(2*1)))" and neuper@37983: d2_abcformula3_neg: neuper@52148: "[|b^^^2 - 4*1*c<0|] ==> (c + b*bdv+ bdv^^^2=0) = False" and neuper@37983: d2_abcformula4: neuper@37906: "[|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: neuper@52148: "[|1 - 4*1*c<0|] ==> (c + bdv+ bdv^^^2=0) = False" and neuper@37983: d2_abcformula5: neuper@37906: "[|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: neuper@52148: "[|Not(bdv occurs_in c); 0 - 4*a*c<0|] ==> (c + a*bdv^^^2=0) = False" and neuper@37983: d2_abcformula6: neuper@37906: "[|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: neuper@52148: "[|Not(bdv occurs_in c); 0 - 4*1*c<0|] ==> (c+ bdv^^^2=0) = False" and neuper@37983: d2_abcformula7: neuper@37906: "[|0<=b^^^2 - 0|] ==> ( b*bdv+a*bdv^^^2=0) = neuper@37906: ((bdv=( -b + sqrt(b^^^2 - 0))/(2*a)) neuper@52148: | (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: neuper@37906: "[|0<=b^^^2 - 0|] ==> ( b*bdv+ bdv^^^2=0) = neuper@37906: ((bdv=( -b + sqrt(b^^^2 - 0))/(2*1)) neuper@52148: | (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: neuper@37906: "[|0<=1 - 0|] ==> ( 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: neuper@37906: "[|0<=1 - 0|] ==> ( 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: neuper@52148: "(a*bdv + b*bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (a + b*bdv + c*bdv^^^2=0))" and neuper@37983: d3_reduce_equation2: neuper@52148: "( bdv + b*bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (1 + b*bdv + c*bdv^^^2=0))" and neuper@37983: d3_reduce_equation3: neuper@52148: "(a*bdv + bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (a + bdv + c*bdv^^^2=0))" and neuper@37983: d3_reduce_equation4: neuper@52148: "( bdv + bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | (1 + bdv + c*bdv^^^2=0))" and neuper@37983: d3_reduce_equation5: neuper@52148: "(a*bdv + b*bdv^^^2 + bdv^^^3=0) = (bdv=0 | (a + b*bdv + bdv^^^2=0))" and neuper@37983: d3_reduce_equation6: neuper@52148: "( bdv + b*bdv^^^2 + bdv^^^3=0) = (bdv=0 | (1 + b*bdv + bdv^^^2=0))" and neuper@37983: d3_reduce_equation7: neuper@52148: "(a*bdv + bdv^^^2 + bdv^^^3=0) = (bdv=0 | (1 + bdv + bdv^^^2=0))" and neuper@37983: d3_reduce_equation8: neuper@52148: "( bdv + bdv^^^2 + bdv^^^3=0) = (bdv=0 | (1 + bdv + bdv^^^2=0))" and neuper@37983: d3_reduce_equation9: neuper@52148: "(a*bdv + c*bdv^^^3=0) = (bdv=0 | (a + c*bdv^^^2=0))" and neuper@37983: d3_reduce_equation10: neuper@52148: "( bdv + c*bdv^^^3=0) = (bdv=0 | (1 + c*bdv^^^2=0))" and neuper@37983: d3_reduce_equation11: neuper@52148: "(a*bdv + bdv^^^3=0) = (bdv=0 | (a + bdv^^^2=0))" and neuper@37983: d3_reduce_equation12: neuper@52148: "( bdv + bdv^^^3=0) = (bdv=0 | (1 + bdv^^^2=0))" and neuper@37983: d3_reduce_equation13: neuper@52148: "( b*bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | ( b*bdv + c*bdv^^^2=0))" and neuper@37983: d3_reduce_equation14: neuper@52148: "( bdv^^^2 + c*bdv^^^3=0) = (bdv=0 | ( bdv + c*bdv^^^2=0))" and neuper@37983: d3_reduce_equation15: neuper@52148: "( b*bdv^^^2 + bdv^^^3=0) = (bdv=0 | ( b*bdv + bdv^^^2=0))" and neuper@37983: d3_reduce_equation16: neuper@52148: "( bdv^^^2 + bdv^^^3=0) = (bdv=0 | ( bdv + bdv^^^2=0))" and neuper@37983: d3_isolate_add1: neuper@52148: "[|Not(bdv occurs_in a)|] ==> (a + b*bdv^^^3=0) = (b*bdv^^^3= (-1)*a)" and neuper@37983: d3_isolate_add2: neuper@52148: "[|Not(bdv occurs_in a)|] ==> (a + bdv^^^3=0) = ( bdv^^^3= (-1)*a)" and neuper@37983: d3_isolate_div: neuper@52148: "[|Not(b=0);Not(bdv occurs_in a)|] ==> (b*bdv^^^3=c) = (bdv^^^3=c/b)" and neuper@37983: d3_root_equation2: neuper@52148: "(bdv^^^3=0) = (bdv=0)" and neuper@37983: d3_root_equation1: neuper@52148: "(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: neuper@37906: "(c+b*bdv^^^2+a*bdv^^^4=0) = neuper@52148: ((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: neuper@52148: bdv_n_collect_1: "l * bdv^^^n + m * bdv^^^n = (l + m) * bdv^^^n" and neuper@52148: bdv_n_collect_2: " bdv^^^n + m * bdv^^^n = (1 + m) * bdv^^^n" and neuper@52148: bdv_n_collect_3: "l * bdv^^^n + bdv^^^n = (l + 1) * bdv^^^n" (*order!*) and neuper@37906: neuper@38030: bdv_n_collect_assoc1_1: neuper@52148: "l * bdv^^^n + (m * bdv^^^n + k) = (l + m) * bdv^^^n + k" and neuper@52148: bdv_n_collect_assoc1_2: "bdv^^^n + (m * bdv^^^n + k) = (1 + m) * bdv^^^n + k" and neuper@52148: bdv_n_collect_assoc1_3: "l * bdv^^^n + (bdv^^^n + k) = (l + 1) * bdv^^^n + k" and neuper@37906: neuper@52148: bdv_n_collect_assoc2_1: "k + l * bdv^^^n + m * bdv^^^n = k +(l + m) * bdv^^^n" and neuper@52148: bdv_n_collect_assoc2_2: "k + bdv^^^n + m * bdv^^^n = k + (1 + m) * bdv^^^n" and neuper@52148: 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 neuper@52148: separate_bdv_n: "(a * bdv ^^^ n) / b = (a / b) * bdv ^^^ n" and neuper@52148: separate_1_bdv: "bdv / b = (1 / b) * (bdv::real)" and neuper@38030: separate_1_bdv_n: "bdv ^^^ n / b = (1 / b) * bdv ^^^ n" neuper@37906: wneuper@59472: ML \ neuper@37972: val thy = @{theory}; neuper@37972: neuper@37954: (*-------------------------rulse-------------------------*) neuper@37954: val PolyEq_prls = (*3.10.02:just the following order due to subterm evaluation*) wneuper@59416: Rule.append_rls "PolyEq_prls" Rule.e_rls wneuper@59416: [Rule.Calc ("Atools.ident",eval_ident "#ident_"), wneuper@59416: Rule.Calc ("Tools.matches",eval_matches ""), wneuper@59416: Rule.Calc ("Tools.lhs" ,eval_lhs ""), wneuper@59416: Rule.Calc ("Tools.rhs" ,eval_rhs ""), wneuper@59416: Rule.Calc ("Poly.is'_expanded'_in",eval_is_expanded_in ""), wneuper@59416: Rule.Calc ("Poly.is'_poly'_in",eval_is_poly_in ""), wneuper@59416: Rule.Calc ("Poly.has'_degree'_in",eval_has_degree_in ""), wneuper@59416: Rule.Calc ("Poly.is'_polyrat'_in",eval_is_polyrat_in ""), wneuper@59416: (*Rule.Calc ("Atools.occurs'_in",eval_occurs_in ""), *) wneuper@59416: (*Rule.Calc ("Atools.is'_const",eval_const "#is_const_"),*) wneuper@59416: Rule.Calc ("HOL.eq",eval_equal "#equal_"), wneuper@59416: Rule.Calc ("RootEq.is'_rootTerm'_in",eval_is_rootTerm_in ""), wneuper@59416: Rule.Calc ("RatEq.is'_ratequation'_in",eval_is_ratequation_in ""), wneuper@59416: Rule.Thm ("not_true",TermC.num_str @{thm not_true}), wneuper@59416: Rule.Thm ("not_false",TermC.num_str @{thm not_false}), wneuper@59416: Rule.Thm ("and_true",TermC.num_str @{thm and_true}), wneuper@59416: Rule.Thm ("and_false",TermC.num_str @{thm and_false}), wneuper@59416: Rule.Thm ("or_true",TermC.num_str @{thm or_true}), wneuper@59416: Rule.Thm ("or_false",TermC.num_str @{thm or_false}) neuper@37954: ]; neuper@37954: neuper@37954: val PolyEq_erls = wneuper@59416: Rule.merge_rls "PolyEq_erls" LinEq_erls wneuper@59416: (Rule.append_rls "ops_preds" calculate_Rational wneuper@59416: [Rule.Calc ("HOL.eq",eval_equal "#equal_"), wneuper@59416: Rule.Thm ("plus_leq", TermC.num_str @{thm plus_leq}), wneuper@59416: Rule.Thm ("minus_leq", TermC.num_str @{thm minus_leq}), wneuper@59416: Rule.Thm ("rat_leq1", TermC.num_str @{thm rat_leq1}), wneuper@59416: Rule.Thm ("rat_leq2", TermC.num_str @{thm rat_leq2}), wneuper@59416: Rule.Thm ("rat_leq3", TermC.num_str @{thm rat_leq3}) neuper@37954: ]); neuper@37954: neuper@37954: val PolyEq_crls = wneuper@59416: Rule.merge_rls "PolyEq_crls" LinEq_crls wneuper@59416: (Rule.append_rls "ops_preds" calculate_Rational wneuper@59416: [Rule.Calc ("HOL.eq",eval_equal "#equal_"), wneuper@59416: Rule.Thm ("plus_leq", TermC.num_str @{thm plus_leq}), wneuper@59416: Rule.Thm ("minus_leq", TermC.num_str @{thm minus_leq}), wneuper@59416: Rule.Thm ("rat_leq1", TermC.num_str @{thm rat_leq1}), wneuper@59416: Rule.Thm ("rat_leq2", TermC.num_str @{thm rat_leq2}), wneuper@59416: Rule.Thm ("rat_leq3", TermC.num_str @{thm rat_leq3}) neuper@37954: ]); neuper@37954: s1210629013@55444: val cancel_leading_coeff = prep_rls'( wneuper@59416: Rule.Rls {id = "cancel_leading_coeff", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), wneuper@59416: erls = PolyEq_erls, srls = Rule.Erls, calc = [], errpatts = [], neuper@37989: rules = wneuper@59416: [Rule.Thm ("cancel_leading_coeff1",TermC.num_str @{thm cancel_leading_coeff1}), wneuper@59416: Rule.Thm ("cancel_leading_coeff2",TermC.num_str @{thm cancel_leading_coeff2}), wneuper@59416: Rule.Thm ("cancel_leading_coeff3",TermC.num_str @{thm cancel_leading_coeff3}), wneuper@59416: Rule.Thm ("cancel_leading_coeff4",TermC.num_str @{thm cancel_leading_coeff4}), wneuper@59416: Rule.Thm ("cancel_leading_coeff5",TermC.num_str @{thm cancel_leading_coeff5}), wneuper@59416: Rule.Thm ("cancel_leading_coeff6",TermC.num_str @{thm cancel_leading_coeff6}), wneuper@59416: Rule.Thm ("cancel_leading_coeff7",TermC.num_str @{thm cancel_leading_coeff7}), wneuper@59416: Rule.Thm ("cancel_leading_coeff8",TermC.num_str @{thm cancel_leading_coeff8}), wneuper@59416: Rule.Thm ("cancel_leading_coeff9",TermC.num_str @{thm cancel_leading_coeff9}), wneuper@59416: Rule.Thm ("cancel_leading_coeff10",TermC.num_str @{thm cancel_leading_coeff10}), wneuper@59416: Rule.Thm ("cancel_leading_coeff11",TermC.num_str @{thm cancel_leading_coeff11}), wneuper@59416: Rule.Thm ("cancel_leading_coeff12",TermC.num_str @{thm cancel_leading_coeff12}), wneuper@59416: Rule.Thm ("cancel_leading_coeff13",TermC.num_str @{thm cancel_leading_coeff13}) wneuper@59416: ],scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")}); s1210629013@55444: wneuper@59374: val prep_rls' = LTool.prep_rls @{theory}; wneuper@59472: \ wneuper@59472: ML\ s1210629013@55444: val complete_square = prep_rls'( wneuper@59416: Rule.Rls {id = "complete_square", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), wneuper@59416: erls = PolyEq_erls, srls = Rule.Erls, calc = [], errpatts = [], wneuper@59416: rules = [Rule.Thm ("complete_square1",TermC.num_str @{thm complete_square1}), wneuper@59416: Rule.Thm ("complete_square2",TermC.num_str @{thm complete_square2}), wneuper@59416: Rule.Thm ("complete_square3",TermC.num_str @{thm complete_square3}), wneuper@59416: Rule.Thm ("complete_square4",TermC.num_str @{thm complete_square4}), wneuper@59416: Rule.Thm ("complete_square5",TermC.num_str @{thm complete_square5}) neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); neuper@37954: s1210629013@55444: val polyeq_simplify = prep_rls'( wneuper@59416: Rule.Rls {id = "polyeq_simplify", preconds = [], neuper@37954: rew_ord = ("termlessI",termlessI), neuper@37954: erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, neuper@42451: calc = [], errpatts = [], wneuper@59416: rules = [Rule.Thm ("real_assoc_1",TermC.num_str @{thm real_assoc_1}), wneuper@59416: Rule.Thm ("real_assoc_2",TermC.num_str @{thm real_assoc_2}), wneuper@59416: Rule.Thm ("real_diff_minus",TermC.num_str @{thm real_diff_minus}), wneuper@59416: Rule.Thm ("real_unari_minus",TermC.num_str @{thm real_unari_minus}), wneuper@59416: Rule.Thm ("realpow_multI",TermC.num_str @{thm realpow_multI}), wneuper@59416: Rule.Calc ("Groups.plus_class.plus",eval_binop "#add_"), wneuper@59416: Rule.Calc ("Groups.minus_class.minus",eval_binop "#sub_"), wneuper@59416: Rule.Calc ("Groups.times_class.times",eval_binop "#mult_"), wneuper@59416: Rule.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"), wneuper@59416: Rule.Calc ("NthRoot.sqrt",eval_sqrt "#sqrt_"), wneuper@59416: Rule.Calc ("Atools.pow" ,eval_binop "#power_"), wneuper@59416: Rule.Rls_ reduce_012 neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); wneuper@59472: \ wneuper@59472: setup \KEStore_Elems.add_rlss neuper@52125: [("cancel_leading_coeff", (Context.theory_name @{theory}, cancel_leading_coeff)), neuper@52125: ("complete_square", (Context.theory_name @{theory}, complete_square)), neuper@52125: ("PolyEq_erls", (Context.theory_name @{theory}, PolyEq_erls)), wneuper@59472: ("polyeq_simplify", (Context.theory_name @{theory}, polyeq_simplify))]\ wneuper@59472: ML\ neuper@37954: neuper@37954: (* ------------- polySolve ------------------ *) 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'( wneuper@59416: Rule.Rls {id = "d0_polyeq_simplify", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), neuper@37954: erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, neuper@42451: calc = [], errpatts = [], wneuper@59416: rules = [Rule.Thm("d0_true",TermC.num_str @{thm d0_true}), wneuper@59416: Rule.Thm("d0_false",TermC.num_str @{thm d0_false}) neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); 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'( wneuper@59416: Rule.Rls {id = "d1_polyeq_simplify", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), neuper@37954: erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, neuper@42451: calc = [], errpatts = [], neuper@37954: rules = [ wneuper@59416: Rule.Thm("d1_isolate_add1",TermC.num_str @{thm d1_isolate_add1}), neuper@37954: (* a+bx=0 -> bx=-a *) wneuper@59416: Rule.Thm("d1_isolate_add2",TermC.num_str @{thm d1_isolate_add2}), neuper@37954: (* a+ x=0 -> x=-a *) wneuper@59416: Rule.Thm("d1_isolate_div",TermC.num_str @{thm d1_isolate_div}) neuper@37954: (* bx=c -> x=c/b *) neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); neuper@37954: wneuper@59472: \ 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'( wneuper@59416: Rule.Rls {id = "d2_polyeq_bdv_only_simplify", preconds = [], rew_ord = ("e_rew_ord",Rule.e_rew_ord), wneuper@59416: erls = PolyEq_erls, srls = Rule.Erls, calc = [], errpatts = [], neuper@42394: rules = wneuper@59416: [Rule.Thm ("d2_prescind1", TermC.num_str @{thm d2_prescind1}), (* ax+bx^2=0 -> x(a+bx)=0 *) wneuper@59416: Rule.Thm ("d2_prescind2", TermC.num_str @{thm d2_prescind2}), (* ax+ x^2=0 -> x(a+ x)=0 *) wneuper@59416: Rule.Thm ("d2_prescind3", TermC.num_str @{thm d2_prescind3}), (* x+bx^2=0 -> x(1+bx)=0 *) wneuper@59416: Rule.Thm ("d2_prescind4", TermC.num_str @{thm d2_prescind4}), (* x+ x^2=0 -> x(1+ x)=0 *) wneuper@59416: Rule.Thm ("d2_sqrt_equation1", TermC.num_str @{thm d2_sqrt_equation1}), (* x^2=c -> x=+-sqrt(c) *) wneuper@59416: Rule.Thm ("d2_sqrt_equation1_neg", TermC.num_str @{thm d2_sqrt_equation1_neg}), (* [0 []*) wneuper@59416: Rule.Thm ("d2_sqrt_equation2", TermC.num_str @{thm d2_sqrt_equation2}), (* x^2=0 -> x=0 *) wneuper@59416: Rule.Thm ("d2_reduce_equation1", TermC.num_str @{thm d2_reduce_equation1}),(* x(a+bx)=0 -> x=0 |a+bx=0*) wneuper@59416: Rule.Thm ("d2_reduce_equation2", TermC.num_str @{thm d2_reduce_equation2}),(* x(a+ x)=0 -> x=0 |a+ x=0*) wneuper@59416: Rule.Thm ("d2_isolate_div", TermC.num_str @{thm d2_isolate_div}) (* bx^2=c -> x^2=c/b *) neuper@42394: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); wneuper@59472: \ wneuper@59472: ML\ neuper@37954: (* 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'( wneuper@59416: Rule.Rls {id = "d2_polyeq_sq_only_simplify", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), neuper@37954: erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, neuper@42451: calc = [], errpatts = [], neuper@37954: (*asm_thm = [("d2_sqrt_equation1",""),("d2_sqrt_equation1_neg",""), neuper@37954: ("d2_isolate_div","")],*) wneuper@59416: rules = [Rule.Thm("d2_isolate_add1",TermC.num_str @{thm d2_isolate_add1}), neuper@37954: (* a+ bx^2=0 -> bx^2=(-1)a*) wneuper@59416: Rule.Thm("d2_isolate_add2",TermC.num_str @{thm d2_isolate_add2}), neuper@37954: (* a+ x^2=0 -> x^2=(-1)a*) wneuper@59416: Rule.Thm("d2_sqrt_equation2",TermC.num_str @{thm d2_sqrt_equation2}), neuper@37954: (* x^2=0 -> x=0 *) wneuper@59416: Rule.Thm("d2_sqrt_equation1",TermC.num_str @{thm d2_sqrt_equation1}), neuper@37954: (* x^2=c -> x=+-sqrt(c)*) wneuper@59416: Rule.Thm("d2_sqrt_equation1_neg",TermC.num_str @{thm d2_sqrt_equation1_neg}), neuper@37954: (* [c<0] x^2=c -> x=[] *) wneuper@59416: Rule.Thm("d2_isolate_div",TermC.num_str @{thm d2_isolate_div}) neuper@37954: (* bx^2=c -> x^2=c/b*) neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); 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'( wneuper@59416: Rule.Rls {id = "d2_polyeq_pqFormula_simplify", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, calc = [], errpatts = [], wneuper@59416: rules = [Rule.Thm("d2_pqformula1",TermC.num_str @{thm d2_pqformula1}), neuper@37954: (* q+px+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula1_neg",TermC.num_str @{thm d2_pqformula1_neg}), neuper@37954: (* q+px+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula2",TermC.num_str @{thm d2_pqformula2}), neuper@37954: (* q+px+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula2_neg",TermC.num_str @{thm d2_pqformula2_neg}), neuper@37954: (* q+px+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula3",TermC.num_str @{thm d2_pqformula3}), neuper@37954: (* q+ x+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula3_neg",TermC.num_str @{thm d2_pqformula3_neg}), neuper@37954: (* q+ x+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula4",TermC.num_str @{thm d2_pqformula4}), neuper@37954: (* q+ x+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula4_neg",TermC.num_str @{thm d2_pqformula4_neg}), neuper@37954: (* q+ x+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula5",TermC.num_str @{thm d2_pqformula5}), neuper@37954: (* qx+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula6",TermC.num_str @{thm d2_pqformula6}), neuper@37954: (* qx+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula7",TermC.num_str @{thm d2_pqformula7}), neuper@37954: (* x+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula8",TermC.num_str @{thm d2_pqformula8}), neuper@37954: (* x+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula9",TermC.num_str @{thm d2_pqformula9}), neuper@37954: (* q +1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula9_neg",TermC.num_str @{thm d2_pqformula9_neg}), neuper@37954: (* q +1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula10",TermC.num_str @{thm d2_pqformula10}), neuper@37954: (* q + x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula10_neg",TermC.num_str @{thm d2_pqformula10_neg}), neuper@37954: (* q + x^2=0 *) wneuper@59416: Rule.Thm("d2_sqrt_equation2",TermC.num_str @{thm d2_sqrt_equation2}), neuper@37954: (* x^2=0 *) wneuper@59416: Rule.Thm("d2_sqrt_equation3",TermC.num_str @{thm d2_sqrt_equation3}) neuper@37954: (* 1x^2=0 *) wneuper@59416: ],scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); wneuper@59472: \ wneuper@59472: ML\ neuper@37954: (* 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'( wneuper@59416: Rule.Rls {id = "d2_polyeq_abcFormula_simplify", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, calc = [], errpatts = [], wneuper@59416: rules = [Rule.Thm("d2_abcformula1",TermC.num_str @{thm d2_abcformula1}), neuper@37954: (*c+bx+cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula1_neg",TermC.num_str @{thm d2_abcformula1_neg}), neuper@37954: (*c+bx+cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula2",TermC.num_str @{thm d2_abcformula2}), neuper@37954: (*c+ x+cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula2_neg",TermC.num_str @{thm d2_abcformula2_neg}), neuper@37954: (*c+ x+cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula3",TermC.num_str @{thm d2_abcformula3}), neuper@37954: (*c+bx+ x^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula3_neg",TermC.num_str @{thm d2_abcformula3_neg}), neuper@37954: (*c+bx+ x^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula4",TermC.num_str @{thm d2_abcformula4}), neuper@37954: (*c+ x+ x^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula4_neg",TermC.num_str @{thm d2_abcformula4_neg}), neuper@37954: (*c+ x+ x^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula5",TermC.num_str @{thm d2_abcformula5}), neuper@37954: (*c+ cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula5_neg",TermC.num_str @{thm d2_abcformula5_neg}), neuper@37954: (*c+ cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula6",TermC.num_str @{thm d2_abcformula6}), neuper@37954: (*c+ x^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula6_neg",TermC.num_str @{thm d2_abcformula6_neg}), neuper@37954: (*c+ x^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula7",TermC.num_str @{thm d2_abcformula7}), neuper@37954: (* bx+ax^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula8",TermC.num_str @{thm d2_abcformula8}), neuper@37954: (* bx+ x^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula9",TermC.num_str @{thm d2_abcformula9}), neuper@37954: (* x+ax^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula10",TermC.num_str @{thm d2_abcformula10}), neuper@37954: (* x+ x^2=0 *) wneuper@59416: Rule.Thm("d2_sqrt_equation2",TermC.num_str @{thm d2_sqrt_equation2}), neuper@37954: (* x^2=0 *) wneuper@59416: Rule.Thm("d2_sqrt_equation3",TermC.num_str @{thm d2_sqrt_equation3}) neuper@37954: (* bx^2=0 *) neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); wneuper@59472: \ wneuper@59472: ML\ 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'( wneuper@59416: Rule.Rls {id = "d2_polyeq_simplify", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, calc = [], errpatts = [], wneuper@59416: rules = [Rule.Thm("d2_pqformula1",TermC.num_str @{thm d2_pqformula1}), neuper@37954: (* p+qx+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula1_neg",TermC.num_str @{thm d2_pqformula1_neg}), neuper@37954: (* p+qx+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula2",TermC.num_str @{thm d2_pqformula2}), neuper@37954: (* p+qx+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula2_neg",TermC.num_str @{thm d2_pqformula2_neg}), neuper@37954: (* p+qx+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula3",TermC.num_str @{thm d2_pqformula3}), neuper@37954: (* p+ x+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula3_neg",TermC.num_str @{thm d2_pqformula3_neg}), neuper@37954: (* p+ x+ x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula4",TermC.num_str @{thm d2_pqformula4}), neuper@37954: (* p+ x+1x^2=0 *) wneuper@59416: Rule.Thm("d2_pqformula4_neg",TermC.num_str @{thm d2_pqformula4_neg}), neuper@37954: (* p+ x+1x^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula1",TermC.num_str @{thm d2_abcformula1}), neuper@37954: (* c+bx+cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula1_neg",TermC.num_str @{thm d2_abcformula1_neg}), neuper@37954: (* c+bx+cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula2",TermC.num_str @{thm d2_abcformula2}), neuper@37954: (* c+ x+cx^2=0 *) wneuper@59416: Rule.Thm("d2_abcformula2_neg",TermC.num_str @{thm d2_abcformula2_neg}), neuper@37954: (* c+ x+cx^2=0 *) wneuper@59416: Rule.Thm("d2_prescind1",TermC.num_str @{thm d2_prescind1}), neuper@37954: (* ax+bx^2=0 -> x(a+bx)=0 *) wneuper@59416: Rule.Thm("d2_prescind2",TermC.num_str @{thm d2_prescind2}), neuper@37954: (* ax+ x^2=0 -> x(a+ x)=0 *) wneuper@59416: Rule.Thm("d2_prescind3",TermC.num_str @{thm d2_prescind3}), neuper@37954: (* x+bx^2=0 -> x(1+bx)=0 *) wneuper@59416: Rule.Thm("d2_prescind4",TermC.num_str @{thm d2_prescind4}), neuper@37954: (* x+ x^2=0 -> x(1+ x)=0 *) wneuper@59416: Rule.Thm("d2_isolate_add1",TermC.num_str @{thm d2_isolate_add1}), neuper@37954: (* a+ bx^2=0 -> bx^2=(-1)a*) wneuper@59416: Rule.Thm("d2_isolate_add2",TermC.num_str @{thm d2_isolate_add2}), neuper@37954: (* a+ x^2=0 -> x^2=(-1)a*) wneuper@59416: Rule.Thm("d2_sqrt_equation1",TermC.num_str @{thm d2_sqrt_equation1}), neuper@37954: (* x^2=c -> x=+-sqrt(c)*) wneuper@59416: Rule.Thm("d2_sqrt_equation1_neg",TermC.num_str @{thm d2_sqrt_equation1_neg}), neuper@37954: (* [c<0] x^2=c -> x=[]*) wneuper@59416: Rule.Thm("d2_sqrt_equation2",TermC.num_str @{thm d2_sqrt_equation2}), neuper@37954: (* x^2=0 -> x=0 *) wneuper@59416: Rule.Thm("d2_reduce_equation1",TermC.num_str @{thm d2_reduce_equation1}), neuper@37954: (* x(a+bx)=0 -> x=0 | a+bx=0*) wneuper@59416: Rule.Thm("d2_reduce_equation2",TermC.num_str @{thm d2_reduce_equation2}), neuper@37954: (* x(a+ x)=0 -> x=0 | a+ x=0*) wneuper@59416: Rule.Thm("d2_isolate_div",TermC.num_str @{thm d2_isolate_div}) neuper@37954: (* bx^2=c -> x^2=c/b*) neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); wneuper@59472: \ wneuper@59472: ML\ 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'( wneuper@59416: Rule.Rls {id = "d3_polyeq_simplify", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, calc = [], errpatts = [], neuper@37954: rules = wneuper@59416: [Rule.Thm("d3_reduce_equation1",TermC.num_str @{thm d3_reduce_equation1}), neuper@37954: (*a*bdv + b*bdv^^^2 + c*bdv^^^3=0) = neuper@37954: (bdv=0 | (a + b*bdv + c*bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation2",TermC.num_str @{thm d3_reduce_equation2}), neuper@37954: (* bdv + b*bdv^^^2 + c*bdv^^^3=0) = neuper@37954: (bdv=0 | (1 + b*bdv + c*bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation3",TermC.num_str @{thm d3_reduce_equation3}), neuper@37954: (*a*bdv + bdv^^^2 + c*bdv^^^3=0) = neuper@37954: (bdv=0 | (a + bdv + c*bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation4",TermC.num_str @{thm d3_reduce_equation4}), neuper@37954: (* bdv + bdv^^^2 + c*bdv^^^3=0) = neuper@37954: (bdv=0 | (1 + bdv + c*bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation5",TermC.num_str @{thm d3_reduce_equation5}), neuper@37954: (*a*bdv + b*bdv^^^2 + bdv^^^3=0) = neuper@37954: (bdv=0 | (a + b*bdv + bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation6",TermC.num_str @{thm d3_reduce_equation6}), neuper@37954: (* bdv + b*bdv^^^2 + bdv^^^3=0) = neuper@37954: (bdv=0 | (1 + b*bdv + bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation7",TermC.num_str @{thm d3_reduce_equation7}), neuper@37954: (*a*bdv + bdv^^^2 + bdv^^^3=0) = neuper@37954: (bdv=0 | (1 + bdv + bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation8",TermC.num_str @{thm d3_reduce_equation8}), neuper@37954: (* bdv + bdv^^^2 + bdv^^^3=0) = neuper@37954: (bdv=0 | (1 + bdv + bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation9",TermC.num_str @{thm d3_reduce_equation9}), neuper@37954: (*a*bdv + c*bdv^^^3=0) = neuper@37954: (bdv=0 | (a + c*bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation10",TermC.num_str @{thm d3_reduce_equation10}), neuper@37954: (* bdv + c*bdv^^^3=0) = neuper@37954: (bdv=0 | (1 + c*bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation11",TermC.num_str @{thm d3_reduce_equation11}), neuper@37954: (*a*bdv + bdv^^^3=0) = neuper@37954: (bdv=0 | (a + bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation12",TermC.num_str @{thm d3_reduce_equation12}), neuper@37954: (* bdv + bdv^^^3=0) = neuper@37954: (bdv=0 | (1 + bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation13",TermC.num_str @{thm d3_reduce_equation13}), neuper@37954: (* b*bdv^^^2 + c*bdv^^^3=0) = neuper@37954: (bdv=0 | ( b*bdv + c*bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation14",TermC.num_str @{thm d3_reduce_equation14}), neuper@37954: (* bdv^^^2 + c*bdv^^^3=0) = neuper@37954: (bdv=0 | ( bdv + c*bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation15",TermC.num_str @{thm d3_reduce_equation15}), neuper@37954: (* b*bdv^^^2 + bdv^^^3=0) = neuper@37954: (bdv=0 | ( b*bdv + bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_reduce_equation16",TermC.num_str @{thm d3_reduce_equation16}), neuper@37954: (* bdv^^^2 + bdv^^^3=0) = neuper@37954: (bdv=0 | ( bdv + bdv^^^2=0)*) wneuper@59416: Rule.Thm("d3_isolate_add1",TermC.num_str @{thm d3_isolate_add1}), neuper@37954: (*[|Not(bdv occurs_in a)|] ==> (a + b*bdv^^^3=0) = neuper@37954: (bdv=0 | (b*bdv^^^3=a)*) wneuper@59416: Rule.Thm("d3_isolate_add2",TermC.num_str @{thm d3_isolate_add2}), neuper@37954: (*[|Not(bdv occurs_in a)|] ==> (a + bdv^^^3=0) = neuper@37954: (bdv=0 | ( bdv^^^3=a)*) wneuper@59416: Rule.Thm("d3_isolate_div",TermC.num_str @{thm d3_isolate_div}), neuper@37954: (*[|Not(b=0)|] ==> (b*bdv^^^3=c) = (bdv^^^3=c/b*) wneuper@59416: Rule.Thm("d3_root_equation2",TermC.num_str @{thm d3_root_equation2}), neuper@37954: (*(bdv^^^3=0) = (bdv=0) *) wneuper@59416: Rule.Thm("d3_root_equation1",TermC.num_str @{thm d3_root_equation1}) neuper@37954: (*bdv^^^3=c) = (bdv = nroot 3 c*) neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); wneuper@59472: \ wneuper@59472: ML\ 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'( wneuper@59416: Rule.Rls {id = "d4_polyeq_simplify", preconds = [], wneuper@59416: rew_ord = ("e_rew_ord",Rule.e_rew_ord), erls = PolyEq_erls, wneuper@59416: srls = Rule.Erls, calc = [], errpatts = [], neuper@37954: rules = wneuper@59416: [Rule.Thm("d4_sub_u1",TermC.num_str @{thm d4_sub_u1}) neuper@37954: (* ax^4+bx^2+c=0 -> x=+-sqrt(ax^2+bx^+c) *) neuper@37954: ], wneuper@59416: scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script") wneuper@59406: }); wneuper@59472: \ wneuper@59472: setup \KEStore_Elems.add_rlss neuper@52125: [("d0_polyeq_simplify", (Context.theory_name @{theory}, d0_polyeq_simplify)), neuper@52125: ("d1_polyeq_simplify", (Context.theory_name @{theory}, d1_polyeq_simplify)), neuper@52125: ("d2_polyeq_simplify", (Context.theory_name @{theory}, d2_polyeq_simplify)), neuper@52125: ("d2_polyeq_bdv_only_simplify", (Context.theory_name @{theory}, d2_polyeq_bdv_only_simplify)), neuper@52125: ("d2_polyeq_sq_only_simplify", (Context.theory_name @{theory}, d2_polyeq_sq_only_simplify)), neuper@52125: neuper@52125: ("d2_polyeq_pqFormula_simplify", neuper@52125: (Context.theory_name @{theory}, d2_polyeq_pqFormula_simplify)), neuper@52125: ("d2_polyeq_abcFormula_simplify", neuper@52125: (Context.theory_name @{theory}, d2_polyeq_abcFormula_simplify)), neuper@52125: ("d3_polyeq_simplify", (Context.theory_name @{theory}, d3_polyeq_simplify)), wneuper@59472: ("d4_polyeq_simplify", (Context.theory_name @{theory}, d4_polyeq_simplify))]\ wneuper@59472: ML\ neuper@37954: (*------------------------problems------------------------*) neuper@37954: (* neuper@37954: (get_pbt ["degree_2","polynomial","univariate","equation"]); neuper@37954: show_ptyps(); neuper@37954: *) wneuper@59472: \ wneuper@59472: setup \KEStore_Elems.add_pbts wneuper@59406: [(Specify.prep_pbt thy "pbl_equ_univ_poly" [] Celem.e_pblID s1210629013@55339: (["polynomial","univariate","equation"], s1210629013@55339: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where" ,["~((e_e::bool) is_ratequation_in (v_v::real))", s1210629013@55339: "~((lhs e_e) is_rootTerm_in (v_v::real))", s1210629013@55339: "~((rhs e_e) is_rootTerm_in (v_v::real))"]), s1210629013@55339: ("#Find" ,["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [])), s1210629013@55339: (*--- d0 ---*) wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg0" [] Celem.e_pblID s1210629013@55339: (["degree_0","polynomial","univariate","equation"], s1210629013@55339: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where" ,["matches (?a = 0) e_e", s1210629013@55339: "(lhs e_e) is_poly_in v_v", s1210629013@55339: "((lhs e_e) has_degree_in v_v ) = 0"]), s1210629013@55339: ("#Find" ,["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [["PolyEq","solve_d0_polyeq_equation"]])), s1210629013@55339: (*--- d1 ---*) wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg1" [] Celem.e_pblID s1210629013@55339: (["degree_1","polynomial","univariate","equation"], s1210629013@55339: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where" ,["matches (?a = 0) e_e", s1210629013@55339: "(lhs e_e) is_poly_in v_v", s1210629013@55339: "((lhs e_e) has_degree_in v_v ) = 1"]), s1210629013@55339: ("#Find" ,["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [["PolyEq","solve_d1_polyeq_equation"]])), s1210629013@55339: (*--- d2 ---*) wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg2" [] Celem.e_pblID s1210629013@55339: (["degree_2","polynomial","univariate","equation"], s1210629013@55339: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where" ,["matches (?a = 0) e_e", s1210629013@55339: "(lhs e_e) is_poly_in v_v ", s1210629013@55339: "((lhs e_e) has_degree_in v_v ) = 2"]), s1210629013@55339: ("#Find" ,["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [["PolyEq","solve_d2_polyeq_equation"]])), wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg2_sqonly" [] Celem.e_pblID s1210629013@55339: (["sq_only","degree_2","polynomial","univariate","equation"], s1210629013@55339: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where" ,["matches ( ?a + ?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches ( ?a + ?b*?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches ( ?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches ( ?b*?v_^^^2 = 0) e_e" , s1210629013@55339: "Not (matches (?a + ?v_ + ?v_^^^2 = 0) e_e) &" ^ s1210629013@55339: "Not (matches (?a + ?b*?v_ + ?v_^^^2 = 0) e_e) &" ^ s1210629013@55339: "Not (matches (?a + ?v_ + ?c*?v_^^^2 = 0) e_e) &" ^ s1210629013@55339: "Not (matches (?a + ?b*?v_ + ?c*?v_^^^2 = 0) e_e) &" ^ s1210629013@55339: "Not (matches ( ?v_ + ?v_^^^2 = 0) e_e) &" ^ s1210629013@55339: "Not (matches ( ?b*?v_ + ?v_^^^2 = 0) e_e) &" ^ s1210629013@55339: "Not (matches ( ?v_ + ?c*?v_^^^2 = 0) e_e) &" ^ s1210629013@55339: "Not (matches ( ?b*?v_ + ?c*?v_^^^2 = 0) e_e)"]), s1210629013@55339: ("#Find" ,["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", s1210629013@55339: [["PolyEq","solve_d2_polyeq_sqonly_equation"]])), wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg2_bdvonly" [] Celem.e_pblID s1210629013@55339: (["bdv_only","degree_2","polynomial","univariate","equation"], s1210629013@55339: [("#Given", ["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where", ["matches (?a*?v_ + ?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches ( ?v_ + ?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches ( ?v_ + ?b*?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches (?a*?v_ + ?b*?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches ( ?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches ( ?b*?v_^^^2 = 0) e_e "]), s1210629013@55339: ("#Find", ["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", s1210629013@55339: [["PolyEq","solve_d2_polyeq_bdvonly_equation"]])), wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg2_pq" [] Celem.e_pblID s1210629013@55339: (["pqFormula","degree_2","polynomial","univariate","equation"], s1210629013@55339: [("#Given", ["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where", ["matches (?a + 1*?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches (?a + ?v_^^^2 = 0) e_e"]), s1210629013@55339: ("#Find", ["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [["PolyEq","solve_d2_polyeq_pq_equation"]])), wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg2_abc" [] Celem.e_pblID s1210629013@55339: (["abcFormula","degree_2","polynomial","univariate","equation"], s1210629013@55339: [("#Given", ["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where", ["matches (?a + ?v_^^^2 = 0) e_e | " ^ s1210629013@55339: "matches (?a + ?b*?v_^^^2 = 0) e_e"]), s1210629013@55339: ("#Find", ["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [["PolyEq","solve_d2_polyeq_abc_equation"]])), s1210629013@55339: (*--- d3 ---*) wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg3" [] Celem.e_pblID s1210629013@55339: (["degree_3","polynomial","univariate","equation"], s1210629013@55339: [("#Given", ["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where", ["matches (?a = 0) e_e", s1210629013@55339: "(lhs e_e) is_poly_in v_v ", s1210629013@55339: "((lhs e_e) has_degree_in v_v) = 3"]), s1210629013@55339: ("#Find", ["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [["PolyEq","solve_d3_polyeq_equation"]])), s1210629013@55339: (*--- d4 ---*) wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_deg4" [] Celem.e_pblID s1210629013@55339: (["degree_4","polynomial","univariate","equation"], s1210629013@55339: [("#Given", ["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where", ["matches (?a = 0) e_e", s1210629013@55339: "(lhs e_e) is_poly_in v_v ", s1210629013@55339: "((lhs e_e) has_degree_in v_v) = 4"]), s1210629013@55339: ("#Find", ["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [(*["PolyEq","solve_d4_polyeq_equation"]*)])), wneuper@59370: (*--- normalise ---*) wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_poly_norm" [] Celem.e_pblID wneuper@59367: (["normalise","polynomial","univariate","equation"], s1210629013@55339: [("#Given", ["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where", ["(Not((matches (?a = 0 ) e_e ))) |" ^ s1210629013@55339: "(Not(((lhs e_e) is_poly_in v_v)))"]), s1210629013@55339: ("#Find", ["solutions v_v'i'"])], wneuper@59370: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [["PolyEq","normalise_poly"]])), s1210629013@55339: (*-------------------------expanded-----------------------*) wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_expand" [] Celem.e_pblID s1210629013@55339: (["expanded","univariate","equation"], s1210629013@55339: [("#Given", ["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where", ["matches (?a = 0) e_e", s1210629013@55339: "(lhs e_e) is_expanded_in v_v "]), s1210629013@55339: ("#Find", ["solutions v_v'i'"])], s1210629013@55339: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [])), s1210629013@55339: (*--- d2 ---*) wneuper@59406: (Specify.prep_pbt thy "pbl_equ_univ_expand_deg2" [] Celem.e_pblID s1210629013@55339: (["degree_2","expanded","univariate","equation"], s1210629013@55339: [("#Given", ["equality e_e","solveFor v_v"]), s1210629013@55339: ("#Where", ["((lhs e_e) has_degree_in v_v) = 2"]), s1210629013@55339: ("#Find", ["solutions v_v'i'"])], wneuper@59472: PolyEq_prls, SOME "solve (e_e::bool, v_v)", [["PolyEq","complete_square"]]))]\ neuper@37954: wneuper@59472: ML\ neuper@37989: val scr = neuper@37989: "Script Normalize_poly (e_e::bool) (v_v::real) = " ^ neuper@37989: "(let e_e =((Try (Rewrite all_left False)) @@ " ^ neuper@37989: " (Try (Repeat (Rewrite makex1_x False))) @@ " ^ neuper@37989: " (Try (Repeat (Rewrite_Set expand_binoms False))) @@ " ^ neuper@37989: " (Try (Repeat (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ neuper@37989: " make_ratpoly_in False))) @@ " ^ neuper@37989: " (Try (Repeat (Rewrite_Set polyeq_simplify False)))) e_e " ^ neuper@37989: " in (SubProblem (PolyEq',[polynomial,univariate,equation], [no_met]) " ^ neuper@37989: " [BOOL e_e, REAL v_v]))"; wneuper@59389: TermC.parse thy scr; wneuper@59472: \ neuper@37954: wneuper@59472: text \"-------------------------methods-----------------------"\ wneuper@59472: setup \KEStore_Elems.add_mets wneuper@59406: [Specify.prep_met thy "met_polyeq" [] Celem.e_metID s1210629013@55373: (["PolyEq"], [], wneuper@59416: {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = Rule.e_rls, prls=Rule.e_rls, s1210629013@55373: crls=PolyEq_crls, errpats = [], nrls = norm_Rational}, s1210629013@55373: "empty_script"), wneuper@59406: Specify.prep_met thy "met_polyeq_norm" [] Celem.e_metID wneuper@59370: (["PolyEq","normalise_poly"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(Not((matches (?a = 0 ) e_e ))) |(Not(((lhs e_e) is_poly_in v_v)))"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule.e_rls, prls=PolyEq_prls, calc=[], s1210629013@55373: crls=PolyEq_crls, errpats = [], nrls = norm_Rational}, s1210629013@55373: "Script Normalize_poly (e_e::bool) (v_v::real) = " ^ s1210629013@55373: "(let e_e =((Try (Rewrite all_left False)) @@ " ^ s1210629013@55373: " (Try (Repeat (Rewrite makex1_x False))) @@ " ^ s1210629013@55373: " (Try (Repeat (Rewrite_Set expand_binoms False))) @@ " ^ s1210629013@55373: " (Try (Repeat (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " make_ratpoly_in False))) @@ " ^ s1210629013@55373: " (Try (Repeat (Rewrite_Set polyeq_simplify False)))) e_e " ^ s1210629013@55373: " in (SubProblem (PolyEq',[polynomial,univariate,equation], [no_met]) " ^ s1210629013@55373: " [BOOL e_e, REAL v_v]))"), wneuper@59406: Specify.prep_met thy "met_polyeq_d0" [] Celem.e_metID s1210629013@55373: (["PolyEq","solve_d0_polyeq_equation"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(lhs e_e) is_poly_in v_v ", "((lhs e_e) has_degree_in v_v) = 0"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule.e_rls, prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Solve_d0_polyeq_equation (e_e::bool) (v_v::real) = " ^ s1210629013@55373: "(let e_e = ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d0_polyeq_simplify False))) e_e " ^ s1210629013@55373: " in ((Or_to_List e_e)::bool list))"), wneuper@59406: Specify.prep_met thy "met_polyeq_d1" [] Celem.e_metID s1210629013@55373: (["PolyEq","solve_d1_polyeq_equation"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(lhs e_e) is_poly_in v_v ", "((lhs e_e) has_degree_in v_v) = 1"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule.e_rls, prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Solve_d1_polyeq_equation (e_e::bool) (v_v::real) = " ^ s1210629013@55373: "(let e_e = ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d1_polyeq_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set norm_Rational_parenthesized False))) e_e;" ^ s1210629013@55373: " (L_L::bool list) = ((Or_to_List e_e)::bool list) " ^ s1210629013@55373: " in Check_elementwise L_L {(v_v::real). Assumptions} )"), wneuper@59406: Specify.prep_met thy "met_polyeq_d22" [] Celem.e_metID s1210629013@55373: (["PolyEq","solve_d2_polyeq_equation"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(lhs e_e) is_poly_in v_v ", "((lhs e_e) has_degree_in v_v) = 2"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule.e_rls, prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Solve_d2_polyeq_equation (e_e::bool) (v_v::real) = " ^ s1210629013@55373: " (let e_e = ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d2_polyeq_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d1_polyeq_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set norm_Rational_parenthesized False))) e_e;" ^ s1210629013@55373: " (L_L::bool list) = ((Or_to_List e_e)::bool list) " ^ s1210629013@55373: " in Check_elementwise L_L {(v_v::real). Assumptions} )"), wneuper@59406: Specify.prep_met thy "met_polyeq_d2_bdvonly" [] Celem.e_metID s1210629013@55373: (["PolyEq","solve_d2_polyeq_bdvonly_equation"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(lhs e_e) is_poly_in v_v ", "((lhs e_e) has_degree_in v_v) = 2"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule.e_rls, prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Solve_d2_polyeq_bdvonly_equation (e_e::bool) (v_v::real) =" ^ s1210629013@55373: " (let e_e = ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d2_polyeq_bdv_only_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d1_polyeq_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set norm_Rational_parenthesized False))) e_e;" ^ s1210629013@55373: " (L_L::bool list) = ((Or_to_List e_e)::bool list) " ^ s1210629013@55373: " in Check_elementwise L_L {(v_v::real). Assumptions} )"), wneuper@59406: Specify.prep_met thy "met_polyeq_d2_sqonly" [] Celem.e_metID s1210629013@55373: (["PolyEq","solve_d2_polyeq_sqonly_equation"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(lhs e_e) is_poly_in v_v ", "((lhs e_e) has_degree_in v_v) = 2"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule.e_rls, prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Solve_d2_polyeq_sqonly_equation (e_e::bool) (v_v::real) =" ^ s1210629013@55373: " (let e_e = ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d2_polyeq_sq_only_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set norm_Rational_parenthesized False))) e_e; " ^ s1210629013@55373: " (L_L::bool list) = ((Or_to_List e_e)::bool list) " ^ s1210629013@55373: " in Check_elementwise L_L {(v_v::real). Assumptions} )"), wneuper@59406: Specify.prep_met thy "met_polyeq_d2_pq" [] Celem.e_metID s1210629013@55373: (["PolyEq","solve_d2_polyeq_pq_equation"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(lhs e_e) is_poly_in v_v ", "((lhs e_e) has_degree_in v_v) = 2"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule.e_rls, prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Solve_d2_polyeq_pq_equation (e_e::bool) (v_v::real) = " ^ s1210629013@55373: " (let e_e = ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d2_polyeq_pqFormula_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set norm_Rational_parenthesized False))) e_e;" ^ s1210629013@55373: " (L_L::bool list) = ((Or_to_List e_e)::bool list) " ^ s1210629013@55373: " in Check_elementwise L_L {(v_v::real). Assumptions} )"), wneuper@59406: Specify.prep_met thy "met_polyeq_d2_abc" [] Celem.e_metID s1210629013@55373: (["PolyEq","solve_d2_polyeq_abc_equation"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(lhs e_e) is_poly_in v_v ", "((lhs e_e) has_degree_in v_v) = 2"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls,srls=Rule.e_rls, prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Solve_d2_polyeq_abc_equation (e_e::bool) (v_v::real) = " ^ s1210629013@55373: " (let e_e = ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d2_polyeq_abcFormula_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set norm_Rational_parenthesized False))) e_e;" ^ s1210629013@55373: " (L_L::bool list) = ((Or_to_List e_e)::bool list) " ^ s1210629013@55373: " in Check_elementwise L_L {(v_v::real). Assumptions} )"), wneuper@59406: Specify.prep_met thy "met_polyeq_d3" [] Celem.e_metID s1210629013@55373: (["PolyEq","solve_d3_polyeq_equation"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["(lhs e_e) is_poly_in v_v ", "((lhs e_e) has_degree_in v_v) = 3"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI", rls'=PolyEq_erls, srls=Rule.e_rls, prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Solve_d3_polyeq_equation (e_e::bool) (v_v::real) = " ^ s1210629013@55373: " (let e_e = ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d3_polyeq_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d2_polyeq_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set_Inst [(bdv,v_v::real)] " ^ s1210629013@55373: " d1_polyeq_simplify True)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set polyeq_simplify False)) @@ " ^ s1210629013@55373: " (Try (Rewrite_Set norm_Rational_parenthesized False))) e_e;" ^ s1210629013@55373: " (L_L::bool list) = ((Or_to_List e_e)::bool list) " ^ s1210629013@55373: " in Check_elementwise L_L {(v_v::real). Assumptions} )"), wneuper@59370: (*.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@59406: Specify.prep_met thy "met_polyeq_complsq" [] Celem.e_metID s1210629013@55373: (["PolyEq","complete_square"], s1210629013@55373: [("#Given" ,["equality e_e","solveFor v_v"]), s1210629013@55373: ("#Where" ,["matches (?a = 0) e_e", "((lhs e_e) has_degree_in v_v) = 2"]), s1210629013@55373: ("#Find" ,["solutions v_v'i'"])], wneuper@59416: {rew_ord'="termlessI",rls'=PolyEq_erls,srls=Rule.e_rls,prls=PolyEq_prls, s1210629013@55373: calc=[("sqrt", ("NthRoot.sqrt", eval_sqrt "#sqrt_"))], crls=PolyEq_crls, errpats = [], s1210629013@55373: nrls = norm_Rational}, s1210629013@55373: "Script Complete_square (e_e::bool) (v_v::real) = " ^ s1210629013@55373: "(let e_e = " ^ s1210629013@55373: " ((Try (Rewrite_Set_Inst [(bdv,v_v)] cancel_leading_coeff True)) " ^ s1210629013@55373: " @@ (Try (Rewrite_Set_Inst [(bdv,v_v)] complete_square True)) " ^ s1210629013@55373: " @@ (Try (Rewrite square_explicit1 False)) " ^ s1210629013@55373: " @@ (Try (Rewrite square_explicit2 False)) " ^ s1210629013@55373: " @@ (Rewrite root_plus_minus True) " ^ s1210629013@55373: " @@ (Try (Repeat (Rewrite_Inst [(bdv,v_v)] bdv_explicit1 False))) " ^ s1210629013@55373: " @@ (Try (Repeat (Rewrite_Inst [(bdv,v_v)] bdv_explicit2 False))) " ^ s1210629013@55373: " @@ (Try (Repeat " ^ s1210629013@55373: " (Rewrite_Inst [(bdv,v_v)] bdv_explicit3 False))) " ^ s1210629013@55373: " @@ (Try (Rewrite_Set calculate_RootRat False)) " ^ s1210629013@55373: " @@ (Try (Repeat (Calculate SQRT)))) e_e " ^ s1210629013@55373: " in ((Or_to_List e_e)::bool list))")] wneuper@59472: \ s1210629013@55373: wneuper@59472: ML\ neuper@37954: neuper@37954: (* termorder hacked by MG *) neuper@37954: 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: neuper@37954: fun dest_hd' x (Const (a, T)) = (((a, 0), T), 0) neuper@37954: | 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) neuper@37954: | dest_hd' x (Var v) = (v, 2) neuper@37954: | dest_hd' x (Bound i) = ((("", i), dummyT), 3) neuper@37954: | dest_hd' x (Abs (_, T, _)) = ((("", 0), T), 4); neuper@37954: neuper@37954: fun size_of_term' x (Const ("Atools.pow",_) $ Free (var,_) $ Free (pot,_)) = neuper@37954: (case x of (*WN*) neuper@37954: (Free (xstr,_)) => wneuper@59390: (if xstr = var then 1000*(the (TermC.int_of_str_opt pot)) else 3) neuper@38031: | _ => error ("size_of_term' called with subst = "^ wneuper@59416: (Rule.term2str x))) neuper@37954: | size_of_term' x (Free (subst,_)) = neuper@37954: (case x of neuper@37954: (Free (xstr,_)) => (if xstr = subst then 1000 else 1) neuper@38031: | _ => error ("size_of_term' called with subst = "^ wneuper@59416: (Rule.term2str x))) neuper@37954: | size_of_term' x (Abs (_,_,body)) = 1 + size_of_term' x body neuper@37954: | size_of_term' x (f$t) = size_of_term' x f + size_of_term' x t neuper@37954: | size_of_term' x _ = 1; neuper@37954: neuper@37989: fun term_ord' x pr thy (Abs (_, T, t), Abs(_, U, u)) = (* ~ term.ML *) neuper@52070: (case term_ord' x pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord) neuper@37989: | term_ord' x pr thy (t, u) = neuper@52070: (if pr neuper@52070: then neuper@52070: let neuper@52070: val (f, ts) = strip_comb t and (g, us) = strip_comb u; wneuper@59416: val _ = tracing ("t= f@ts= \"" ^ Rule.term_to_string''' thy f ^ "\" @ \"[" ^ wneuper@59416: commas (map (Rule.term_to_string''' thy) ts) ^ "]\""); wneuper@59416: val _ = tracing ("u= g@us= \"" ^ Rule.term_to_string''' thy g ^ "\" @ \"[" ^ wneuper@59416: commas(map (Rule.term_to_string''' thy) us) ^ "]\""); neuper@52070: val _ = tracing ("size_of_term(t,u)= (" ^ string_of_int (size_of_term' x t) ^ ", " ^ neuper@52070: string_of_int (size_of_term' x u) ^ ")"); neuper@52070: val _ = tracing ("hd_ord(f,g) = " ^ (pr_ord o (hd_ord x)) (f,g)); neuper@52070: val _ = tracing ("terms_ord(ts,us) = " ^ (pr_ord o (terms_ord x) str false) (ts, us)); neuper@52070: val _ = tracing ("-------"); neuper@52070: in () end neuper@52070: else (); neuper@52070: case int_ord (size_of_term' x t, size_of_term' x u) of neuper@52070: EQUAL => neuper@52070: let val (f, ts) = strip_comb t and (g, us) = strip_comb u neuper@52070: in neuper@52070: (case hd_ord x (f, g) of neuper@52070: EQUAL => (terms_ord x str pr) (ts, us) neuper@52070: | ord => ord) neuper@52070: end neuper@37954: | ord => ord) neuper@37954: and hd_ord x (f, g) = (* ~ term.ML *) neuper@37989: prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) neuper@37989: int_ord (dest_hd' x f, dest_hd' x g) neuper@37954: and terms_ord x str pr (ts, us) = wneuper@59406: list_ord (term_ord' x pr (Celem.assoc_thy "Isac"))(ts, us); neuper@52070: neuper@37954: in neuper@37954: neuper@37954: fun ord_make_polynomial_in (pr:bool) thy subst tu = neuper@37954: let wneuper@59406: (* val _=tracing("*** subs variable is: "^(Celem.subst2str subst)); *) neuper@37954: in neuper@37954: case subst of neuper@37954: (_,x)::_ => (term_ord' x pr thy tu = LESS) neuper@38031: | _ => error ("ord_make_polynomial_in called with subst = "^ wneuper@59406: (Celem.subst2str subst)) neuper@37954: end; neuper@37989: end;(*local*) neuper@37954: wneuper@59472: \ wneuper@59472: ML\ s1210629013@55444: val order_add_mult_in = prep_rls'( wneuper@59416: Rule.Rls{id = "order_add_mult_in", preconds = [], neuper@37954: rew_ord = ("ord_make_polynomial_in", neuper@52139: ord_make_polynomial_in false @{theory "Poly"}), wneuper@59416: erls = Rule.e_rls,srls = Rule.Erls, neuper@42451: calc = [], errpatts = [], wneuper@59416: rules = [Rule.Thm ("mult_commute",TermC.num_str @{thm mult.commute}), neuper@37954: (* z * w = w * z *) wneuper@59416: Rule.Thm ("real_mult_left_commute",TermC.num_str @{thm real_mult_left_commute}), neuper@37954: (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*) wneuper@59416: Rule.Thm ("mult_assoc",TermC.num_str @{thm mult.assoc}), neuper@37954: (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*) wneuper@59416: Rule.Thm ("add_commute",TermC.num_str @{thm add.commute}), neuper@37954: (*z + w = w + z*) wneuper@59416: Rule.Thm ("add_left_commute",TermC.num_str @{thm add.left_commute}), neuper@37954: (*x + (y + z) = y + (x + z)*) wneuper@59416: Rule.Thm ("add_assoc",TermC.num_str @{thm add.assoc}) neuper@37954: (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*) wneuper@59416: ], scr = Rule.EmptyScr}); neuper@37954: wneuper@59472: \ wneuper@59472: ML\ s1210629013@55444: val collect_bdv = prep_rls'( wneuper@59416: Rule.Rls{id = "collect_bdv", preconds = [], wneuper@59416: rew_ord = ("dummy_ord", Rule.dummy_ord), wneuper@59416: erls = Rule.e_rls,srls = Rule.Erls, neuper@42451: calc = [], errpatts = [], wneuper@59416: rules = [Rule.Thm ("bdv_collect_1",TermC.num_str @{thm bdv_collect_1}), wneuper@59416: Rule.Thm ("bdv_collect_2",TermC.num_str @{thm bdv_collect_2}), wneuper@59416: Rule.Thm ("bdv_collect_3",TermC.num_str @{thm bdv_collect_3}), neuper@37954: wneuper@59416: Rule.Thm ("bdv_collect_assoc1_1",TermC.num_str @{thm bdv_collect_assoc1_1}), wneuper@59416: Rule.Thm ("bdv_collect_assoc1_2",TermC.num_str @{thm bdv_collect_assoc1_2}), wneuper@59416: Rule.Thm ("bdv_collect_assoc1_3",TermC.num_str @{thm bdv_collect_assoc1_3}), neuper@37954: wneuper@59416: Rule.Thm ("bdv_collect_assoc2_1",TermC.num_str @{thm bdv_collect_assoc2_1}), wneuper@59416: Rule.Thm ("bdv_collect_assoc2_2",TermC.num_str @{thm bdv_collect_assoc2_2}), wneuper@59416: Rule.Thm ("bdv_collect_assoc2_3",TermC.num_str @{thm bdv_collect_assoc2_3}), neuper@37954: neuper@37954: wneuper@59416: Rule.Thm ("bdv_n_collect_1",TermC.num_str @{thm bdv_n_collect_1}), wneuper@59416: Rule.Thm ("bdv_n_collect_2",TermC.num_str @{thm bdv_n_collect_2}), wneuper@59416: Rule.Thm ("bdv_n_collect_3",TermC.num_str @{thm bdv_n_collect_3}), neuper@37954: wneuper@59416: Rule.Thm ("bdv_n_collect_assoc1_1",TermC.num_str @{thm bdv_n_collect_assoc1_1}), wneuper@59416: Rule.Thm ("bdv_n_collect_assoc1_2",TermC.num_str @{thm bdv_n_collect_assoc1_2}), wneuper@59416: Rule.Thm ("bdv_n_collect_assoc1_3",TermC.num_str @{thm bdv_n_collect_assoc1_3}), neuper@37954: wneuper@59416: Rule.Thm ("bdv_n_collect_assoc2_1",TermC.num_str @{thm bdv_n_collect_assoc2_1}), wneuper@59416: Rule.Thm ("bdv_n_collect_assoc2_2",TermC.num_str @{thm bdv_n_collect_assoc2_2}), wneuper@59416: Rule.Thm ("bdv_n_collect_assoc2_3",TermC.num_str @{thm bdv_n_collect_assoc2_3}) wneuper@59416: ], scr = Rule.EmptyScr}); 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'( wneuper@59416: Rule.Seq {id = "make_polynomial_in", preconds = []:term list, wneuper@59416: rew_ord = ("dummy_ord", Rule.dummy_ord), wneuper@59416: erls = Atools_erls, srls = Rule.Erls, neuper@42451: calc = [], errpatts = [], wneuper@59416: rules = [Rule.Rls_ expand_poly, wneuper@59416: Rule.Rls_ order_add_mult_in, wneuper@59416: Rule.Rls_ simplify_power, wneuper@59416: Rule.Rls_ collect_numerals, wneuper@59416: Rule.Rls_ reduce_012, wneuper@59416: Rule.Thm ("realpow_oneI",TermC.num_str @{thm realpow_oneI}), wneuper@59416: Rule.Rls_ discard_parentheses, wneuper@59416: Rule.Rls_ collect_bdv neuper@37954: ], wneuper@59416: scr = Rule.EmptyScr wneuper@59406: }); neuper@37954: wneuper@59472: \ wneuper@59472: ML\ neuper@37954: val separate_bdvs = wneuper@59416: Rule.append_rls "separate_bdvs" neuper@37954: collect_bdv wneuper@59416: [Rule.Thm ("separate_bdv", TermC.num_str @{thm separate_bdv}), neuper@37954: (*"?a * ?bdv / ?b = ?a / ?b * ?bdv"*) wneuper@59416: Rule.Thm ("separate_bdv_n", TermC.num_str @{thm separate_bdv_n}), wneuper@59416: Rule.Thm ("separate_1_bdv", TermC.num_str @{thm separate_1_bdv}), neuper@37954: (*"?bdv / ?b = (1 / ?b) * ?bdv"*) wneuper@59416: Rule.Thm ("separate_1_bdv_n", TermC.num_str @{thm separate_1_bdv_n}), neuper@37954: (*"?bdv ^^^ ?n / ?b = 1 / ?b * ?bdv ^^^ ?n"*) wneuper@59416: Rule.Thm ("add_divide_distrib", wneuper@59389: TermC.num_str @{thm add_divide_distrib}) neuper@37954: (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z" neuper@37954: WN051031 DOES NOT BELONG TO HERE*) neuper@37954: ]; wneuper@59472: \ wneuper@59472: ML\ s1210629013@55444: val make_ratpoly_in = prep_rls'( wneuper@59416: Rule.Seq {id = "make_ratpoly_in", preconds = []:term list, wneuper@59416: rew_ord = ("dummy_ord", Rule.dummy_ord), wneuper@59416: erls = Atools_erls, srls = Rule.Erls, neuper@42451: calc = [], errpatts = [], wneuper@59416: rules = [Rule.Rls_ norm_Rational, wneuper@59416: Rule.Rls_ order_add_mult_in, wneuper@59416: Rule.Rls_ discard_parentheses, wneuper@59416: Rule.Rls_ separate_bdvs, wneuper@59416: (* Rule.Rls_ rearrange_assoc, WN060916 why does cancel_p not work?*) wneuper@59416: Rule.Rls_ cancel_p wneuper@59416: (*Rule.Calc ("Rings.divide_class.divide" ,eval_cancel "#divide_e") too weak!*) neuper@37954: ], wneuper@59416: scr = Rule.EmptyScr}); wneuper@59472: \ wneuper@59472: setup \KEStore_Elems.add_rlss neuper@52130: [("order_add_mult_in", (Context.theory_name @{theory}, order_add_mult_in)), neuper@52130: ("collect_bdv", (Context.theory_name @{theory}, collect_bdv)), neuper@52130: ("make_polynomial_in", (Context.theory_name @{theory}, make_polynomial_in)), neuper@52130: ("make_ratpoly_in", (Context.theory_name @{theory}, make_ratpoly_in)), wneuper@59472: ("separate_bdvs", (Context.theory_name @{theory}, separate_bdvs))]\ neuper@37954: neuper@37906: end neuper@37906: neuper@37906: neuper@37906: neuper@37906: neuper@37906: neuper@37906: