src/Tools/isac/Knowledge/LinEq.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 04 Aug 2022 12:48:37 +0200
changeset 60509 2e0b7ca391dc
parent 60449 2406d378cede
child 60515 03e19793d81e
permissions -rw-r--r--
polish naming in Rewrite_Order
     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 axiomatization where
    13 (*-- normalise --*)
    14   (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*)
    15   all_left:          "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)" and
    16   makex1_x:          "a \<up> 1  = a" and
    17   real_assoc_1:      "a+(b+c) = a+b+c" and
    18   real_assoc_2:      "a*(b*c) = a*b*c" and
    19 
    20 (*-- solve --*)
    21   lin_isolate_add1:  "(a + b*bdv = 0) = (b*bdv = (-1)*a)" and
    22   lin_isolate_add2:  "(a +   bdv = 0) = (  bdv = (-1)*a)" and
    23   lin_isolate_div:   "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)"
    24 
    25 ML \<open>
    26 val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
    27   Rule_Set.append_rules "LinEq_prls" Rule_Set.empty [
    28     \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
    29     \<^rule_eval>\<open>Prog_Expr.matches\<close> (Prog_Expr.eval_matches ""),
    30     \<^rule_eval>\<open>Prog_Expr.lhs\<close> (Prog_Expr.eval_lhs ""),
    31     \<^rule_eval>\<open>Prog_Expr.rhs\<close> (Prog_Expr.eval_rhs ""),
    32     \<^rule_eval>\<open>has_degree_in\<close> (eval_has_degree_in ""),
    33     \<^rule_eval>\<open>is_polyrat_in\<close> (eval_is_polyrat_in ""),
    34     \<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in ""),    
    35     \<^rule_eval>\<open>Prog_Expr.ident\<close> (Prog_Expr.eval_ident "#ident_"),
    36     \<^rule_thm>\<open>not_true\<close>,
    37     \<^rule_thm>\<open>not_false\<close>,
    38     \<^rule_thm>\<open>and_true\<close>,
    39     \<^rule_thm>\<open>and_false\<close>,
    40     \<^rule_thm>\<open>or_true\<close>,
    41     \<^rule_thm>\<open>or_false\<close>];
    42 (* ----- erls ----- *)
    43 val LinEq_crls = 
    44    Rule_Set.append_rules "LinEq_crls" poly_crls
    45    [\<^rule_thm>\<open>real_assoc_1\<close>
    46     (*		
    47      Don't use
    48      \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
    49      \<^rule_eval>\<open>realpow\<close> (**)(eval_binop "#power_"),
    50      *)
    51     ];
    52 
    53 (* ----- crls ----- *)
    54 val LinEq_erls = 
    55   Rule_Set.append_rules "LinEq_erls" Poly_erls [
    56     \<^rule_thm>\<open>real_assoc_1\<close>
    57     (*		
    58      Don't use
    59      \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
    60      \<^rule_eval>\<open>realpow\<close> (**)(eval_binop "#power_"),
    61      *)
    62     ];
    63 \<close>
    64 rule_set_knowledge LinEq_erls = LinEq_erls
    65 ML \<open>
    66     
    67 val LinPoly_simplify = prep_rls'(
    68   Rule_Def.Repeat {id = "LinPoly_simplify", preconds = [], 
    69     rew_ord = ("termlessI",termlessI), 
    70     erls = LinEq_erls, 
    71     srls = Rule_Set.Empty, 
    72     calc = [], errpatts = [],
    73     rules = [
    74       \<^rule_thm>\<open>real_assoc_1\<close>,
    75       \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
    76       \<^rule_eval>\<open>minus\<close> (**)(eval_binop "#sub_"),
    77       \<^rule_eval>\<open>times\<close> (**)(eval_binop "#mult_"),
    78       (*  Don't use  
    79        \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
    80        \<^rule_eval>\<open>sqrt\<close> (eval_sqrt "#sqrt_"),
    81        *)
    82       \<^rule_eval>\<open>realpow\<close> (**)(eval_binop "#power_")],
    83     scr = Rule.Empty_Prog});
    84 \<close>
    85 rule_set_knowledge LinPoly_simplify = LinPoly_simplify
    86 ML \<open>
    87 
    88 (*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*)
    89 val LinEq_simplify = prep_rls'(
    90   Rule_Def.Repeat {id = "LinEq_simplify", preconds = [],
    91     rew_ord = ("Rewrite_Ord.id_empty", Rewrite_Ord.function_empty),
    92     erls = LinEq_erls,
    93     srls = Rule_Set.Empty,
    94     calc = [], errpatts = [],
    95     rules = [
    96 	     \<^rule_thm>\<open>lin_isolate_add1\<close>, (* a+bx=0 -> bx=-a *)
    97 	     \<^rule_thm>\<open>lin_isolate_add2\<close>, (* a+ x=0 ->  x=-a *)
    98 	     \<^rule_thm>\<open>lin_isolate_div\<close> (*   bx=c -> x=c/b *)],
    99     scr = Rule.Empty_Prog});
   100 \<close>
   101 rule_set_knowledge LinEq_simplify = LinEq_simplify
   102 
   103 (*----------------------------- problems --------------------------------*)
   104 (* ---------linear----------- *)
   105 
   106 problem pbl_equ_univ_lin : "LINEAR/univariate/equation" =
   107   \<open>LinEq_prls\<close>
   108   Method_Ref: "LinEq/solve_lineq_equation"
   109   CAS: "solve (e_e::bool, v_v)"
   110   Given: "equality e_e" "solveFor v_v"
   111   Where: "HOL.False" (*WN0509 just detected: this pbl can never be used?!?*)
   112     "Not( (lhs e_e) is_polyrat_in v_v)"
   113     "Not( (rhs e_e) is_polyrat_in v_v)"
   114     "((lhs e_e) has_degree_in v_v)=1"
   115 	  "((rhs e_e) has_degree_in v_v)=1"
   116   Find: "solutions v_v'i'"
   117 
   118 (*-------------- methods------------------------------------------------------*)
   119 method met_eqlin : "LinEq" =
   120   \<open>{rew_ord' = "tless_true",rls' = Atools_erls,calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty,
   121     crls = LinEq_crls, errpats = [], nrls = norm_Poly}\<close>
   122     (* ansprechen mit ["LinEq", "solve_univar_equation"] *)
   123 
   124 partial_function (tailrec) solve_linear_equation :: "bool \<Rightarrow> real \<Rightarrow> bool list"
   125   where
   126 "solve_linear_equation e_e v_v = (
   127   let
   128     e_e = (
   129       (Try (Rewrite ''all_left'')) #>
   130       (Try (Repeat (Rewrite ''makex1_x''))) #>
   131       (Try (Rewrite_Set ''expand_binoms'')) #>
   132       (Try (Repeat (Rewrite_Set_Inst [(''bdv'', v_v)] ''make_ratpoly_in''))) #>
   133       (Try (Repeat (Rewrite_Set ''LinPoly_simplify''))) ) e_e;
   134     e_e = (
   135       (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''LinEq_simplify'')) #>
   136       (Repeat (Try (Rewrite_Set ''LinPoly_simplify''))) ) e_e
   137   in
   138     Or_to_List e_e)"
   139 
   140 method met_eq_lin : "LinEq/solve_lineq_equation" =
   141   \<open>{rew_ord' = "termlessI", rls' = LinEq_erls, srls = Rule_Set.empty, prls = LinEq_prls, calc = [],
   142     crls = LinEq_crls, errpats = [], nrls = norm_Poly}\<close>
   143   Program: solve_linear_equation.simps
   144   Given: "equality e_e" "solveFor v_v"
   145   Where: "Not ((lhs e_e) is_polyrat_in v_v)" "((lhs e_e)  has_degree_in v_v) = 1"
   146   Find: "solutions v_v'i'"
   147 
   148 ML \<open>
   149   MethodC.from_store' @{theory} ["LinEq", "solve_lineq_equation"];
   150 \<close> ML \<open>
   151 \<close> ML \<open>
   152 \<close>
   153 
   154 end
   155