neuper@37906: (*. (c) by Richard Lang, 2003 .*) neuper@37906: (* theory collecting all knowledge for LinearEquations neuper@37906: created by: rlang neuper@37906: date: 02.10 neuper@37906: changed by: rlang neuper@37906: last change by: rlang neuper@37906: date: 02.10.20 neuper@37906: *) neuper@37906: neuper@37950: theory LinEq imports Poly Equation begin neuper@37906: neuper@37906: consts neuper@37906: Solve'_lineq'_equation neuper@37950: :: "[bool,real, neuper@37950: bool list] => bool list" wneuper@59334: ("((Script Solve'_lineq'_equation (_ _ =))// (_))" 9) neuper@37906: neuper@52148: axiomatization where wneuper@59370: (*-- normalise --*) neuper@37906: (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*) neuper@52148: all_left: "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*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: (*-- solve --*) neuper@52148: lin_isolate_add1: "(a + b*bdv = 0) = (b*bdv = (-1)*a)" and neuper@52148: lin_isolate_add2: "(a + bdv = 0) = ( bdv = (-1)*a)" and neuper@37982: lin_isolate_div: "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)" neuper@37950: neuper@37950: ML {* neuper@37972: val thy = @{theory}; neuper@37972: neuper@37950: val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*) wneuper@59406: Celem.append_rls "LinEq_prls" Celem.e_rls wneuper@59406: [Celem.Calc ("HOL.eq",eval_equal "#equal_"), wneuper@59406: Celem.Calc ("Tools.matches",eval_matches ""), wneuper@59406: Celem.Calc ("Tools.lhs" ,eval_lhs ""), wneuper@59406: Celem.Calc ("Tools.rhs" ,eval_rhs ""), wneuper@59406: Celem.Calc ("Poly.has'_degree'_in",eval_has_degree_in ""), wneuper@59406: Celem.Calc ("Poly.is'_polyrat'_in",eval_is_polyrat_in ""), wneuper@59406: Celem.Calc ("Atools.occurs'_in",eval_occurs_in ""), wneuper@59406: Celem.Calc ("Atools.ident",eval_ident "#ident_"), wneuper@59406: Celem.Thm ("not_true",TermC.num_str @{thm not_true}), wneuper@59406: Celem.Thm ("not_false",TermC.num_str @{thm not_false}), wneuper@59406: Celem.Thm ("and_true",TermC.num_str @{thm and_true}), wneuper@59406: Celem.Thm ("and_false",TermC.num_str @{thm and_false}), wneuper@59406: Celem.Thm ("or_true",TermC.num_str @{thm or_true}), wneuper@59406: Celem.Thm ("or_false",TermC.num_str @{thm or_false}) neuper@37950: ]; neuper@37950: (* ----- erls ----- *) neuper@37950: val LinEq_crls = wneuper@59406: Celem.append_rls "LinEq_crls" poly_crls wneuper@59406: [Celem.Thm ("real_assoc_1",TermC.num_str @{thm real_assoc_1}) neuper@37950: (* neuper@37950: Don't use wneuper@59406: Celem.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"), wneuper@59406: Celem.Calc ("Atools.pow" ,eval_binop "#power_"), neuper@37950: *) neuper@37950: ]; neuper@37950: neuper@37950: (* ----- crls ----- *) neuper@37950: val LinEq_erls = wneuper@59406: Celem.append_rls "LinEq_erls" Poly_erls wneuper@59406: [Celem.Thm ("real_assoc_1",TermC.num_str @{thm real_assoc_1}) neuper@37950: (* neuper@37950: Don't use wneuper@59406: Celem.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"), wneuper@59406: Celem.Calc ("Atools.pow" ,eval_binop "#power_"), neuper@37950: *) neuper@37950: ]; neuper@52125: *} neuper@52125: setup {* KEStore_Elems.add_rlss neuper@52125: [("LinEq_erls", (Context.theory_name @{theory}, LinEq_erls))] *} neuper@52125: ML {* neuper@37950: s1210629013@55444: val LinPoly_simplify = prep_rls'( wneuper@59406: Celem.Rls {id = "LinPoly_simplify", preconds = [], neuper@37950: rew_ord = ("termlessI",termlessI), neuper@37950: erls = LinEq_erls, wneuper@59406: srls = Celem.Erls, neuper@42451: calc = [], errpatts = [], neuper@37950: rules = [ wneuper@59406: Celem.Thm ("real_assoc_1",TermC.num_str @{thm real_assoc_1}), wneuper@59406: Celem.Calc ("Groups.plus_class.plus",eval_binop "#add_"), wneuper@59406: Celem.Calc ("Groups.minus_class.minus",eval_binop "#sub_"), wneuper@59406: Celem.Calc ("Groups.times_class.times",eval_binop "#mult_"), neuper@37950: (* Dont use wneuper@59406: Celem.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"), wneuper@59406: Celem.Calc ("NthRoot.sqrt",eval_sqrt "#sqrt_"), neuper@37950: *) wneuper@59406: Celem.Calc ("Atools.pow" ,eval_binop "#power_") neuper@37950: ], wneuper@59406: scr = Celem.EmptyScr}); neuper@52125: *} neuper@52125: setup {* KEStore_Elems.add_rlss neuper@52125: [("LinPoly_simplify", (Context.theory_name @{theory}, LinPoly_simplify))] *} neuper@52125: ML {* neuper@37950: neuper@37950: (*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*) s1210629013@55444: val LinEq_simplify = prep_rls'( wneuper@59406: Celem.Rls {id = "LinEq_simplify", preconds = [], wneuper@59411: rew_ord = ("e_rew_ord", Celem.e_rew_ord), neuper@37950: erls = LinEq_erls, wneuper@59406: srls = Celem.Erls, neuper@42451: calc = [], errpatts = [], neuper@37950: rules = [ wneuper@59406: Celem.Thm("lin_isolate_add1",TermC.num_str @{thm lin_isolate_add1}), neuper@37950: (* a+bx=0 -> bx=-a *) wneuper@59406: Celem.Thm("lin_isolate_add2",TermC.num_str @{thm lin_isolate_add2}), neuper@37950: (* a+ x=0 -> x=-a *) wneuper@59406: Celem.Thm("lin_isolate_div",TermC.num_str @{thm lin_isolate_div}) neuper@37950: (* bx=c -> x=c/b *) neuper@37950: ], wneuper@59406: scr = Celem.EmptyScr}); neuper@52125: *} neuper@52125: setup {* KEStore_Elems.add_rlss neuper@52125: [("LinEq_simplify", (Context.theory_name @{theory}, LinEq_simplify))] *} neuper@37950: neuper@37950: (*----------------------------- problem types --------------------------------*) neuper@37950: (* ---------linear----------- *) s1210629013@55359: setup {* KEStore_Elems.add_pbts wneuper@59406: [(Specify.prep_pbt thy "pbl_equ_univ_lin" [] Celem.e_pblID s1210629013@55339: (["LINEAR", "univariate", "equation"], s1210629013@55339: [("#Given" ,["equality e_e", "solveFor v_v"]), s1210629013@55339: ("#Where" ,["HOL.False", (*WN0509 just detected: this pbl can never be used?!?*) s1210629013@55339: "Not( (lhs e_e) is_polyrat_in v_v)", s1210629013@55339: "Not( (rhs e_e) is_polyrat_in v_v)", s1210629013@55339: "((lhs e_e) has_degree_in v_v)=1", s1210629013@55339: "((rhs e_e) has_degree_in v_v)=1"]), s1210629013@55339: ("#Find" ,["solutions v_v'i'"])], s1210629013@55339: LinEq_prls, SOME "solve (e_e::bool, v_v)", [["LinEq", "solve_lineq_equation"]]))] *} neuper@37950: s1210629013@55373: (*-------------- methods------------------------------------------------------*) s1210629013@55373: setup {* KEStore_Elems.add_mets wneuper@59406: [Specify.prep_met thy "met_eqlin" [] Celem.e_metID s1210629013@55373: (["LinEq"], [], wneuper@59406: {rew_ord' = "tless_true",rls' = Atools_erls,calc = [], srls = Celem.e_rls, prls = Celem.e_rls, s1210629013@55373: crls = LinEq_crls, errpats = [], nrls = norm_Poly}, s1210629013@55373: "empty_script"), s1210629013@55373: (* ansprechen mit ["LinEq","solve_univar_equation"] *) wneuper@59406: Specify.prep_met thy "met_eq_lin" [] Celem.e_metID s1210629013@55373: (["LinEq","solve_lineq_equation"], s1210629013@55373: [("#Given", ["equality e_e", "solveFor v_v"]), s1210629013@55373: ("#Where", ["Not ((lhs e_e) is_polyrat_in v_v)", "((lhs e_e) has_degree_in v_v) = 1"]), s1210629013@55373: ("#Find", ["solutions v_v'i'"])], wneuper@59406: {rew_ord' = "termlessI", rls' = LinEq_erls, srls = Celem.e_rls, prls = LinEq_prls, calc = [], s1210629013@55373: crls = LinEq_crls, errpats = [], nrls = norm_Poly}, s1210629013@55373: "Script Solve_lineq_equation (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 (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 LinPoly_simplify False))))e_e;" ^ s1210629013@55373: " e_e = ((Try (Rewrite_Set_Inst [(bdv, v_v::real)] " ^ s1210629013@55373: " LinEq_simplify True)) @@ " ^ s1210629013@55373: " (Repeat(Try (Rewrite_Set LinPoly_simplify False)))) e_e " ^ s1210629013@55373: " in ((Or_to_List e_e)::bool list))")] s1210629013@55373: *} wneuper@59269: ML {* Specify.get_met' @{theory} ["LinEq","solve_lineq_equation"]; *} neuper@37950: neuper@37906: end neuper@37906: