src/Tools/isac/Knowledge/LinEq.thy
author wenzelm
Wed, 26 May 2021 16:24:05 +0200
changeset 60286 31efa1b39a20
parent 60278 343efa173023
child 60289 a7b88fc19a92
permissions -rw-r--r--
command 'setup_rule' semantic equivalent for KEStore_Elems.add_rlss;
     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 thy = @{theory};
    27 
    28 val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
    29   Rule_Set.append_rules "LinEq_prls" Rule_Set.empty 
    30 	     [Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
    31 	      Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches ""),
    32 	      Rule.Eval ("Prog_Expr.lhs"    , Prog_Expr.eval_lhs ""),
    33 	      Rule.Eval ("Prog_Expr.rhs"    , Prog_Expr.eval_rhs ""),
    34 	      Rule.Eval ("Poly.has_degree_in", eval_has_degree_in ""),
    35  	      Rule.Eval ("Poly.is_polyrat_in", eval_is_polyrat_in ""),
    36 	      Rule.Eval ("Prog_Expr.occurs_in", Prog_Expr.eval_occurs_in ""),    
    37 	      Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),
    38 	      Rule.Thm ("not_true",ThmC.numerals_to_Free @{thm not_true}),
    39 	      Rule.Thm ("not_false",ThmC.numerals_to_Free @{thm not_false}),
    40 	      Rule.Thm ("and_true",ThmC.numerals_to_Free @{thm and_true}),
    41 	      Rule.Thm ("and_false",ThmC.numerals_to_Free @{thm and_false}),
    42 	      Rule.Thm ("or_true",ThmC.numerals_to_Free @{thm or_true}),
    43 	      Rule.Thm ("or_false",ThmC.numerals_to_Free @{thm or_false})
    44               ];
    45 (* ----- erls ----- *)
    46 val LinEq_crls = 
    47    Rule_Set.append_rules "LinEq_crls" poly_crls
    48    [Rule.Thm  ("real_assoc_1",ThmC.numerals_to_Free @{thm real_assoc_1})
    49     (*		
    50      Don't use
    51      Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),
    52      Rule.Eval ("Transcendental.powr" , (**)eval_binop "#power_"),
    53      *)
    54     ];
    55 
    56 (* ----- crls ----- *)
    57 val LinEq_erls = 
    58    Rule_Set.append_rules "LinEq_erls" Poly_erls
    59    [Rule.Thm  ("real_assoc_1",ThmC.numerals_to_Free @{thm real_assoc_1})
    60     (*		
    61      Don't use
    62      Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),
    63      Rule.Eval ("Transcendental.powr" , (**)eval_binop "#power_"),
    64      *)
    65     ];
    66 \<close>
    67 setup_rule LinEq_erls = LinEq_erls
    68 ML \<open>
    69     
    70 val LinPoly_simplify = prep_rls'(
    71   Rule_Def.Repeat {id = "LinPoly_simplify", preconds = [], 
    72        rew_ord = ("termlessI",termlessI), 
    73        erls = LinEq_erls, 
    74        srls = Rule_Set.Empty, 
    75        calc = [], errpatts = [],
    76        rules = [
    77 		Rule.Thm  ("real_assoc_1",ThmC.numerals_to_Free @{thm real_assoc_1}),
    78 		Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_"),
    79 		Rule.Eval ("Groups.minus_class.minus", (**)eval_binop "#sub_"),
    80 		Rule.Eval ("Groups.times_class.times", (**)eval_binop "#mult_"),
    81 		(*  Dont use  
    82 		 Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),		
    83 		 Rule.Eval ("NthRoot.sqrt", eval_sqrt "#sqrt_"),
    84 		 *)
    85 		Rule.Eval ("Transcendental.powr" , (**)eval_binop "#power_")
    86 		],
    87        scr = Rule.Empty_Prog});
    88 \<close>
    89 setup_rule LinPoly_simplify = LinPoly_simplify
    90 ML \<open>
    91 
    92 (*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*)
    93 val LinEq_simplify = prep_rls'(
    94 Rule_Def.Repeat {id = "LinEq_simplify", preconds = [],
    95      rew_ord = ("e_rew_ord", Rewrite_Ord.e_rew_ord),
    96      erls = LinEq_erls,
    97      srls = Rule_Set.Empty,
    98      calc = [], errpatts = [],
    99      rules = [
   100 	      Rule.Thm("lin_isolate_add1",ThmC.numerals_to_Free @{thm lin_isolate_add1}), 
   101 	      (* a+bx=0 -> bx=-a *)
   102 	      Rule.Thm("lin_isolate_add2",ThmC.numerals_to_Free @{thm lin_isolate_add2}), 
   103 	      (* a+ x=0 ->  x=-a *)
   104 	      Rule.Thm("lin_isolate_div",ThmC.numerals_to_Free @{thm lin_isolate_div})    
   105 	      (*   bx=c -> x=c/b *)  
   106 	      ],
   107      scr = Rule.Empty_Prog});
   108 \<close>
   109 setup_rule LinEq_simplify = LinEq_simplify
   110 
   111 (*----------------------------- problem types --------------------------------*)
   112 (* ---------linear----------- *)
   113 setup \<open>KEStore_Elems.add_pbts
   114   [(Problem.prep_input thy "pbl_equ_univ_lin" [] Problem.id_empty
   115       (["LINEAR", "univariate", "equation"],
   116         [("#Given" ,["equality e_e", "solveFor v_v"]),
   117           ("#Where" ,["HOL.False", (*WN0509 just detected: this pbl can never be used?!?*)
   118               "Not( (lhs e_e) is_polyrat_in v_v)",
   119               "Not( (rhs e_e) is_polyrat_in v_v)",
   120               "((lhs e_e) has_degree_in v_v)=1",
   121 	            "((rhs e_e) has_degree_in v_v)=1"]),
   122           ("#Find"  ,["solutions v_v'i'"])],
   123         LinEq_prls, SOME "solve (e_e::bool, v_v)", [["LinEq", "solve_lineq_equation"]]))]\<close>
   124 
   125 (*-------------- methods------------------------------------------------------*)
   126 setup \<open>KEStore_Elems.add_mets
   127     [MethodC.prep_input thy "met_eqlin" [] MethodC.id_empty
   128       (["LinEq"], [],
   129         {rew_ord' = "tless_true",rls' = Atools_erls,calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty,
   130           crls = LinEq_crls, errpats = [], nrls = norm_Poly},
   131         @{thm refl})]
   132 \<close>
   133     (* ansprechen mit ["LinEq", "solve_univar_equation"] *)
   134 
   135 partial_function (tailrec) solve_linear_equation :: "bool \<Rightarrow> real \<Rightarrow> bool list"
   136   where
   137 "solve_linear_equation e_e v_v = (
   138   let
   139     e_e = (
   140       (Try (Rewrite ''all_left'')) #>
   141       (Try (Repeat (Rewrite ''makex1_x''))) #>
   142       (Try (Rewrite_Set ''expand_binoms'')) #>
   143       (Try (Repeat (Rewrite_Set_Inst [(''bdv'', v_v)] ''make_ratpoly_in''))) #>
   144       (Try (Repeat (Rewrite_Set ''LinPoly_simplify''))) ) e_e;
   145     e_e = (
   146       (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''LinEq_simplify'')) #>
   147       (Repeat (Try (Rewrite_Set ''LinPoly_simplify''))) ) e_e
   148   in
   149     Or_to_List e_e)"
   150 setup \<open>KEStore_Elems.add_mets
   151     [MethodC.prep_input thy "met_eq_lin" [] MethodC.id_empty
   152       (["LinEq", "solve_lineq_equation"],
   153         [("#Given", ["equality e_e", "solveFor v_v"]),
   154           ("#Where", ["Not ((lhs e_e) is_polyrat_in v_v)", "((lhs e_e)  has_degree_in v_v) = 1"]),
   155           ("#Find",  ["solutions v_v'i'"])],
   156         {rew_ord' = "termlessI", rls' = LinEq_erls, srls = Rule_Set.empty, prls = LinEq_prls, calc = [],
   157           crls = LinEq_crls, errpats = [], nrls = norm_Poly},
   158         @{thm solve_linear_equation.simps})]
   159 \<close>
   160 ML \<open>
   161   MethodC.from_store' @{theory} ["LinEq", "solve_lineq_equation"];
   162 \<close> ML \<open>
   163 \<close> ML \<open>
   164 \<close>
   165 
   166 end
   167