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