src/Tools/isac/Knowledge/LinEq.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 28 Nov 2018 11:46:00 +0100
changeset 59473 28b67cae58c3
parent 59472 3e904f8ec16c
child 59489 cfcbcac0bae8
permissions -rw-r--r--
funpack: separate programs in prep. for partial_function
     1 (*. (c) by Richard Lang, 2003 .*)
     2 (* theory collecting all knowledge for LinearEquations
     3    created by: rlang 
     4          date: 02.10
     5    changed by: rlang
     6    last change by: rlang
     7              date: 02.10.20
     8 *)
     9 
    10 theory LinEq imports Poly Equation begin
    11 
    12 consts
    13    Solve'_lineq'_equation
    14              :: "[bool,real, 
    15 		   bool list] => bool list"
    16                ("((Script Solve'_lineq'_equation (_ _ =))// (_))" 9)
    17 
    18 axiomatization where
    19 (*-- normalise --*)
    20   (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*)
    21   all_left:          "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)" and
    22   makex1_x:          "a^^^1  = a" and
    23   real_assoc_1:      "a+(b+c) = a+b+c" and
    24   real_assoc_2:      "a*(b*c) = a*b*c" and
    25 
    26 (*-- solve --*)
    27   lin_isolate_add1:  "(a + b*bdv = 0) = (b*bdv = (-1)*a)" and
    28   lin_isolate_add2:  "(a +   bdv = 0) = (  bdv = (-1)*a)" and
    29   lin_isolate_div:   "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)"
    30 
    31 ML \<open>
    32 val thy = @{theory};
    33 
    34 val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
    35   Rule.append_rls "LinEq_prls" Rule.e_rls 
    36 	     [Rule.Calc ("HOL.eq",eval_equal "#equal_"),
    37 	      Rule.Calc ("Tools.matches",eval_matches ""),
    38 	      Rule.Calc ("Tools.lhs"    ,eval_lhs ""),
    39 	      Rule.Calc ("Tools.rhs"    ,eval_rhs ""),
    40 	      Rule.Calc ("Poly.has'_degree'_in",eval_has_degree_in ""),
    41  	      Rule.Calc ("Poly.is'_polyrat'_in",eval_is_polyrat_in ""),
    42 	      Rule.Calc ("Atools.occurs'_in",eval_occurs_in ""),    
    43 	      Rule.Calc ("Atools.ident",eval_ident "#ident_"),
    44 	      Rule.Thm ("not_true",TermC.num_str @{thm not_true}),
    45 	      Rule.Thm ("not_false",TermC.num_str @{thm not_false}),
    46 	      Rule.Thm ("and_true",TermC.num_str @{thm and_true}),
    47 	      Rule.Thm ("and_false",TermC.num_str @{thm and_false}),
    48 	      Rule.Thm ("or_true",TermC.num_str @{thm or_true}),
    49 	      Rule.Thm ("or_false",TermC.num_str @{thm or_false})
    50               ];
    51 (* ----- erls ----- *)
    52 val LinEq_crls = 
    53    Rule.append_rls "LinEq_crls" poly_crls
    54    [Rule.Thm  ("real_assoc_1",TermC.num_str @{thm real_assoc_1})
    55     (*		
    56      Don't use
    57      Rule.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),
    58      Rule.Calc ("Atools.pow" ,eval_binop "#power_"),
    59      *)
    60     ];
    61 
    62 (* ----- crls ----- *)
    63 val LinEq_erls = 
    64    Rule.append_rls "LinEq_erls" Poly_erls
    65    [Rule.Thm  ("real_assoc_1",TermC.num_str @{thm real_assoc_1})
    66     (*		
    67      Don't use
    68      Rule.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),
    69      Rule.Calc ("Atools.pow" ,eval_binop "#power_"),
    70      *)
    71     ];
    72 \<close>
    73 setup \<open>KEStore_Elems.add_rlss 
    74   [("LinEq_erls", (Context.theory_name @{theory}, LinEq_erls))]\<close>
    75 ML \<open>
    76     
    77 val LinPoly_simplify = prep_rls'(
    78   Rule.Rls {id = "LinPoly_simplify", preconds = [], 
    79        rew_ord = ("termlessI",termlessI), 
    80        erls = LinEq_erls, 
    81        srls = Rule.Erls, 
    82        calc = [], errpatts = [],
    83        rules = [
    84 		Rule.Thm  ("real_assoc_1",TermC.num_str @{thm real_assoc_1}),
    85 		Rule.Calc ("Groups.plus_class.plus",eval_binop "#add_"),
    86 		Rule.Calc ("Groups.minus_class.minus",eval_binop "#sub_"),
    87 		Rule.Calc ("Groups.times_class.times",eval_binop "#mult_"),
    88 		(*  Dont use  
    89 		 Rule.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),		
    90 		 Rule.Calc ("NthRoot.sqrt",eval_sqrt "#sqrt_"),
    91 		 *)
    92 		Rule.Calc ("Atools.pow" ,eval_binop "#power_")
    93 		],
    94        scr = Rule.EmptyScr});
    95 \<close>
    96 setup \<open>KEStore_Elems.add_rlss 
    97   [("LinPoly_simplify", (Context.theory_name @{theory}, LinPoly_simplify))]\<close>
    98 ML \<open>
    99 
   100 (*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*)
   101 val LinEq_simplify = prep_rls'(
   102 Rule.Rls {id = "LinEq_simplify", preconds = [],
   103      rew_ord = ("e_rew_ord", Rule.e_rew_ord),
   104      erls = LinEq_erls,
   105      srls = Rule.Erls,
   106      calc = [], errpatts = [],
   107      rules = [
   108 	      Rule.Thm("lin_isolate_add1",TermC.num_str @{thm lin_isolate_add1}), 
   109 	      (* a+bx=0 -> bx=-a *)
   110 	      Rule.Thm("lin_isolate_add2",TermC.num_str @{thm lin_isolate_add2}), 
   111 	      (* a+ x=0 ->  x=-a *)
   112 	      Rule.Thm("lin_isolate_div",TermC.num_str @{thm lin_isolate_div})    
   113 	      (*   bx=c -> x=c/b *)  
   114 	      ],
   115      scr = Rule.EmptyScr});
   116 \<close>
   117 setup \<open>KEStore_Elems.add_rlss 
   118   [("LinEq_simplify", (Context.theory_name @{theory}, LinEq_simplify))]\<close>
   119 
   120 (*----------------------------- problem types --------------------------------*)
   121 (* ---------linear----------- *)
   122 setup \<open>KEStore_Elems.add_pbts
   123   [(Specify.prep_pbt thy "pbl_equ_univ_lin" [] Celem.e_pblID
   124       (["LINEAR", "univariate", "equation"],
   125         [("#Given" ,["equality e_e", "solveFor v_v"]),
   126           ("#Where" ,["HOL.False", (*WN0509 just detected: this pbl can never be used?!?*)
   127               "Not( (lhs e_e) is_polyrat_in v_v)",
   128               "Not( (rhs e_e) is_polyrat_in v_v)",
   129               "((lhs e_e) has_degree_in v_v)=1",
   130 	            "((rhs e_e) has_degree_in v_v)=1"]),
   131           ("#Find"  ,["solutions v_v'i'"])],
   132         LinEq_prls, SOME "solve (e_e::bool, v_v)", [["LinEq", "solve_lineq_equation"]]))]\<close>
   133 
   134 (*-------------- methods------------------------------------------------------*)
   135 setup \<open>KEStore_Elems.add_mets
   136     [Specify.prep_met thy "met_eqlin" [] Celem.e_metID
   137       (["LinEq"], [],
   138         {rew_ord' = "tless_true",rls' = Atools_erls,calc = [], srls = Rule.e_rls, prls = Rule.e_rls,
   139           crls = LinEq_crls, errpats = [], nrls = norm_Poly},
   140         "empty_script")]
   141 \<close>
   142     (* ansprechen mit ["LinEq","solve_univar_equation"] *)
   143 setup \<open>KEStore_Elems.add_mets
   144     [Specify.prep_met thy "met_eq_lin" [] Celem.e_metID
   145       (["LinEq","solve_lineq_equation"],
   146         [("#Given", ["equality e_e", "solveFor v_v"]),
   147           ("#Where", ["Not ((lhs e_e) is_polyrat_in v_v)", "((lhs e_e)  has_degree_in v_v) = 1"]),
   148           ("#Find",  ["solutions v_v'i'"])],
   149         {rew_ord' = "termlessI", rls' = LinEq_erls, srls = Rule.e_rls, prls = LinEq_prls, calc = [],
   150           crls = LinEq_crls, errpats = [], nrls = norm_Poly},
   151         "Script Solve_lineq_equation (e_e::bool) (v_v::real) =                 " ^
   152           "(let e_e =((Try         (Rewrite      all_left           False)) @@   " ^ 
   153           "           (Try (Repeat (Rewrite     makex1_x            False))) @@  " ^ 
   154           "           (Try         (Rewrite_Set expand_binoms       False)) @@   " ^ 
   155           "           (Try (Repeat (Rewrite_Set_Inst [(bdv, v_v::real)]          " ^
   156           "                                 make_ratpoly_in    False)))    @@    " ^
   157           "           (Try (Repeat (Rewrite_Set LinPoly_simplify      False))))e_e;" ^
   158           "     e_e = ((Try (Rewrite_Set_Inst [(bdv, v_v::real)]                  " ^
   159           "                                     LinEq_simplify True)) @@  " ^
   160           "            (Repeat(Try (Rewrite_Set LinPoly_simplify     False)))) e_e " ^
   161           " in ((Or_to_List e_e)::bool list))")]
   162 \<close>
   163 ML \<open>Specify.get_met' @{theory} ["LinEq","solve_lineq_equation"];\<close>
   164 
   165 end
   166