src/Tools/isac/Knowledge/LinEq.thy
author Walther Neuper <neuper@ist.tugraz.at>
Tue, 28 Sep 2010 10:10:26 +0200
branchisac-update-Isa09-2
changeset 38034 928cebc9c4aa
parent 38014 3e11e3c2dc42
child 41922 32d7766945fb
permissions -rw-r--r--
updated "op *" --> Groups.times_class.times in src and test

find . -type f -exec sed -i s/"\"op \*\""/"\"Groups.times_class.times\""/g {} \;
     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 (_ _ =))// 
    17                   (_))" 9)
    18 
    19 axioms
    20 (*-- normalize --*)
    21   (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*)
    22   all_left:          "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)"
    23   makex1_x:          "a^^^1  = a"  
    24   real_assoc_1:      "a+(b+c) = a+b+c"
    25   real_assoc_2:      "a*(b*c) = a*b*c"
    26 
    27 (*-- solve --*)
    28   lin_isolate_add1:  "(a + b*bdv = 0) = (b*bdv = (-1)*a)"
    29   lin_isolate_add2:  "(a +   bdv = 0) = (  bdv = (-1)*a)"
    30   lin_isolate_div:   "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)"
    31 
    32 ML {*
    33 val thy = @{theory};
    34 
    35 val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
    36   append_rls "LinEq_prls" e_rls 
    37 	     [Calc ("op =",eval_equal "#equal_"),
    38 	      Calc ("Tools.matches",eval_matches ""),
    39 	      Calc ("Tools.lhs"    ,eval_lhs ""),
    40 	      Calc ("Tools.rhs"    ,eval_rhs ""),
    41 	      Calc ("Poly.has'_degree'_in",eval_has_degree_in ""),
    42  	      Calc ("Poly.is'_polyrat'_in",eval_is_polyrat_in ""),
    43 	      Calc ("Atools.occurs'_in",eval_occurs_in ""),    
    44 	      Calc ("Atools.ident",eval_ident "#ident_"),
    45 	      Thm ("not_true",num_str @{thm not_true}),
    46 	      Thm ("not_false",num_str @{thm not_false}),
    47 	      Thm ("and_true",num_str @{thm and_true}),
    48 	      Thm ("and_false",num_str @{thm and_false}),
    49 	      Thm ("or_true",num_str @{thm or_true}),
    50 	      Thm ("or_false",num_str @{thm or_false})
    51               ];
    52 (* ----- erls ----- *)
    53 val LinEq_crls = 
    54    append_rls "LinEq_crls" poly_crls
    55    [Thm  ("real_assoc_1",num_str @{thm real_assoc_1})
    56     (*		
    57      Don't use
    58      Calc ("Rings.inverse_class.divide", eval_cancel "#divide_e"),
    59      Calc ("Atools.pow" ,eval_binop "#power_"),
    60      *)
    61     ];
    62 
    63 (* ----- crls ----- *)
    64 val LinEq_erls = 
    65    append_rls "LinEq_erls" Poly_erls
    66    [Thm  ("real_assoc_1",num_str @{thm real_assoc_1})
    67     (*		
    68      Don't use
    69      Calc ("Rings.inverse_class.divide", eval_cancel "#divide_e"),
    70      Calc ("Atools.pow" ,eval_binop "#power_"),
    71      *)
    72     ];
    73 
    74 ruleset' := overwritelthy @{theory} (!ruleset',
    75 			[("LinEq_erls",LinEq_erls)(*FIXXXME:del with rls.rls'*)
    76 			 ]);
    77     
    78 val LinPoly_simplify = prep_rls(
    79   Rls {id = "LinPoly_simplify", preconds = [], 
    80        rew_ord = ("termlessI",termlessI), 
    81        erls = LinEq_erls, 
    82        srls = Erls, 
    83        calc = [], 
    84        (*asm_thm = [],*)
    85        rules = [
    86 		Thm  ("real_assoc_1",num_str @{thm real_assoc_1}),
    87 		Calc ("Groups.plus_class.plus",eval_binop "#add_"),
    88 		Calc ("Groups.minus_class.minus",eval_binop "#sub_"),
    89 		Calc ("Groups.times_class.times",eval_binop "#mult_"),
    90 		(*  Dont use  
    91 		 Calc ("Rings.inverse_class.divide", eval_cancel "#divide_e"),		
    92 		 Calc ("NthRoot.sqrt",eval_sqrt "#sqrt_"),
    93 		 *)
    94 		Calc ("Atools.pow" ,eval_binop "#power_")
    95 		],
    96        scr = EmptyScr}:rls);
    97 
    98 ruleset' := overwritelthy @{theory} (!ruleset',
    99 			  [("LinPoly_simplify",LinPoly_simplify)]);
   100 
   101 (*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*)
   102 val LinEq_simplify = prep_rls(
   103 Rls {id = "LinEq_simplify", preconds = [],
   104      rew_ord = ("e_rew_ord",e_rew_ord),
   105      erls = LinEq_erls,
   106      srls = Erls,
   107      calc = [],
   108      (*asm_thm = [("lin_isolate_div","")],*)
   109      rules = [
   110 	      Thm("lin_isolate_add1",num_str @{thm lin_isolate_add1}), 
   111 	      (* a+bx=0 -> bx=-a *)
   112 	      Thm("lin_isolate_add2",num_str @{thm lin_isolate_add2}), 
   113 	      (* a+ x=0 ->  x=-a *)
   114 	      Thm("lin_isolate_div",num_str @{thm lin_isolate_div})    
   115 	      (*   bx=c -> x=c/b *)  
   116 	      ],
   117      scr = EmptyScr}:rls);
   118 ruleset' := overwritelthy @{theory} (!ruleset',
   119 			[("LinEq_simplify",LinEq_simplify)]);
   120 
   121 (*----------------------------- problem types --------------------------------*)
   122 (* 
   123 show_ptyps(); 
   124 (get_pbt ["linear","univariate","equation"]);
   125 *)
   126 
   127 (* ---------linear----------- *)
   128 store_pbt
   129  (prep_pbt thy "pbl_equ_univ_lin" [] e_pblID
   130  (["linear","univariate","equation"],
   131   [("#Given" ,["equality e_e","solveFor v_v"]),
   132    ("#Where" ,["False", (*WN0509 just detected: this pbl can never be used?!?*)
   133                "Not( (lhs e_e) is_polyrat_in v_v)",
   134                "Not( (rhs e_e) is_polyrat_in v_v)",
   135                "((lhs e_e) has_degree_in v_v)=1",
   136 	       "((rhs e_e) has_degree_in v_v)=1"]),
   137    ("#Find"  ,["solutions v_v'i'"]) 
   138   ],
   139   LinEq_prls, SOME "solve (e_e::bool, v_v)",
   140   [["LinEq","solve_lineq_equation"]]));
   141 
   142 (*-------------- methods------------------------------------------------------*)
   143 store_met
   144  (prep_met thy "met_eqlin" [] e_metID
   145  (["LinEq"],
   146    [],
   147    {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = e_rls, prls=e_rls,
   148     crls=LinEq_crls, nrls=norm_Poly
   149     (*, asm_rls=[],asm_thm=[]*)}, "empty_script"));
   150 
   151 (* ansprechen mit ["LinEq","solve_univar_equation"] *)
   152 store_met
   153 (prep_met thy "met_eq_lin" [] e_metID
   154  (["LinEq","solve_lineq_equation"],
   155    [("#Given", ["equality e_e", "solveFor v_v"]),
   156     ("#Where", ["Not ((lhs e_e) is_polyrat_in v_v)",
   157                 "((lhs e_e)  has_degree_in v_v) = 1"]),
   158     ("#Find",  ["solutions v_v'i'"])
   159    ],
   160    {rew_ord'="termlessI", rls'=LinEq_erls, srls=e_rls, prls=LinEq_prls,
   161     calc=[], crls=LinEq_crls, nrls=norm_Poly},
   162     "Script Solve_lineq_equation (e_e::bool) (v_v::real) =                 " ^
   163     "(let e_e =((Try         (Rewrite      all_left           False)) @@   " ^ 
   164     "           (Try (Repeat (Rewrite     makex1_x            False))) @@  " ^ 
   165     "           (Try         (Rewrite_Set expand_binoms       False)) @@   " ^ 
   166     "           (Try (Repeat (Rewrite_Set_Inst [(bdv, v_v::real)]          " ^
   167     "                                 make_ratpoly_in    False)))    @@    " ^
   168     "           (Try (Repeat (Rewrite_Set LinPoly_simplify      False))))e_e;" ^
   169     "     e_e = ((Try (Rewrite_Set_Inst [(bdv, v_v::real)]                  " ^
   170     "                                     LinEq_simplify True)) @@  " ^
   171     "            (Repeat(Try (Rewrite_Set LinPoly_simplify     False)))) e_e " ^
   172     " in ((Or_to_List e_e)::bool list))"
   173  ));
   174 get_met ["LinEq","solve_lineq_equation"];
   175 
   176 *}
   177 
   178 end
   179