src/Tools/isac/Knowledge/LinEq.thy
author wneuper <Walther.Neuper@jku.at>
Mon, 01 Jan 2024 11:31:16 +0100
changeset 60789 8fa678b678e8
parent 60624 0e0ac7706f0d
permissions -rw-r--r--
Doc/Specify_Phase 4: start use antiquotations from isar-ref
     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 section \<open>Theorems for Rule_Set; FIXME: replace axiomatization\<close>
    13 
    14 axiomatization where
    15 (*-- normalise --*)
    16   (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*)
    17   all_left:          "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)" and
    18   makex1_x:          "a \<up> 1  = a" and
    19   real_assoc_1:      "a+(b+c) = a+b+c" and
    20   real_assoc_2:      "a*(b*c) = a*b*c" and
    21 
    22 (*-- solve --*)
    23   lin_isolate_add1:  "(a + b*bdv = 0) = (b*bdv = (-1)*a)" and
    24   lin_isolate_add2:  "(a +   bdv = 0) = (  bdv = (-1)*a)" and
    25   lin_isolate_div:   "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)"
    26 
    27 
    28 section \<open>Rule_Set for Problem and MethodC\<close>
    29 
    30 ML \<open>
    31 val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
    32   Rule_Set.append_rules "LinEq_prls" Rule_Set.empty [
    33     \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
    34     \<^rule_eval>\<open>Prog_Expr.matches\<close> (Prog_Expr.eval_matches ""),
    35     \<^rule_eval>\<open>Prog_Expr.lhs\<close> (Prog_Expr.eval_lhs ""),
    36     \<^rule_eval>\<open>Prog_Expr.rhs\<close> (Prog_Expr.eval_rhs ""),
    37     \<^rule_eval>\<open>has_degree_in\<close> (eval_has_degree_in ""),
    38     \<^rule_eval>\<open>is_polyrat_in\<close> (eval_is_polyrat_in ""),
    39     \<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in ""),    
    40     \<^rule_eval>\<open>Prog_Expr.ident\<close> (Prog_Expr.eval_ident "#ident_"),
    41     \<^rule_thm>\<open>not_true\<close>,
    42     \<^rule_thm>\<open>not_false\<close>,
    43     \<^rule_thm>\<open>and_true\<close>,
    44     \<^rule_thm>\<open>and_false\<close>,
    45     \<^rule_thm>\<open>or_true\<close>,
    46     \<^rule_thm>\<open>or_false\<close>];
    47 
    48 (* ----- asm_rls ----- *)
    49 val LinEq_crls = 
    50    Rule_Set.append_rules "LinEq_crls" poly_crls
    51    [\<^rule_thm>\<open>real_assoc_1\<close>
    52     (*		
    53      Don't use
    54      \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
    55      \<^rule_eval>\<open>realpow\<close> (**)(Calc_Binop.numeric "#power_"),
    56      *)
    57     ];
    58 
    59 (* ----- crls ----- *)
    60 val LinEq_erls = 
    61   Rule_Set.append_rules "LinEq_erls" Poly_erls [
    62     \<^rule_thm>\<open>real_assoc_1\<close>
    63     (*		
    64      Don't use
    65      \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
    66      \<^rule_eval>\<open>realpow\<close> (**)(Calc_Binop.numeric "#power_"),
    67      *)
    68     ];
    69 \<close>
    70 rule_set_knowledge LinEq_erls = LinEq_erls
    71 ML \<open>
    72     
    73 val LinPoly_simplify = prep_rls'(
    74   Rule_Def.Repeat {id = "LinPoly_simplify", preconds = [], 
    75     rew_ord = ("termlessI",termlessI), 
    76     asm_rls = LinEq_erls, 
    77     prog_rls = Rule_Set.Empty, 
    78     calc = [], errpatts = [],
    79     rules = [
    80       \<^rule_thm>\<open>real_assoc_1\<close>,
    81       \<^rule_eval>\<open>plus\<close> (**)(Calc_Binop.numeric "#add_"),
    82       \<^rule_eval>\<open>minus\<close> (**)(Calc_Binop.numeric "#sub_"),
    83       \<^rule_eval>\<open>times\<close> (**)(Calc_Binop.numeric "#mult_"),
    84       (*  Don't use  
    85        \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
    86        \<^rule_eval>\<open>sqrt\<close> (eval_sqrt "#sqrt_"),
    87        *)
    88       \<^rule_eval>\<open>realpow\<close> (**)(Calc_Binop.numeric "#power_")],
    89     program = Rule.Empty_Prog});
    90 \<close>
    91 rule_set_knowledge LinPoly_simplify = LinPoly_simplify
    92 
    93 ML \<open>
    94 (*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*)
    95 val LinEq_simplify = prep_rls'(
    96   Rule_Def.Repeat {id = "LinEq_simplify", preconds = [],
    97     rew_ord = ("Rewrite_Ord.id_empty", Rewrite_Ord.function_empty),
    98     asm_rls = LinEq_erls,
    99     prog_rls = Rule_Set.Empty,
   100     calc = [], errpatts = [],
   101     rules = [
   102 	     \<^rule_thm>\<open>lin_isolate_add1\<close>, (* a+bx=0 -> bx=-a *)
   103 	     \<^rule_thm>\<open>lin_isolate_add2\<close>, (* a+ x=0 ->  x=-a *)
   104 	     \<^rule_thm>\<open>lin_isolate_div\<close> (*   bx=c -> x=c/b *)],
   105     program = Rule.Empty_Prog});
   106 \<close>
   107 rule_set_knowledge LinEq_simplify = LinEq_simplify
   108 
   109 
   110 section \<open>Problem.T\<close>
   111 
   112 problem pbl_equ_univ_lin : "LINEAR/univariate/equation" =
   113   \<open>LinEq_prls\<close>
   114   Method_Ref: "LinEq/solve_lineq_equation"
   115   CAS: "solve (e_e::bool, v_v)"
   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 
   124 
   125 section \<open>MethodC.T\<close>
   126 
   127 method met_eqlin : "LinEq" =
   128   \<open>{rew_ord = "tless_true",rls' = Atools_erls,calc = [], prog_rls = Rule_Set.empty, where_rls = Rule_Set.empty,
   129     errpats = [], rew_rls = norm_Poly}\<close>
   130     (* ansprechen mit ["LinEq", "solve_univar_equation"] *)
   131 
   132 partial_function (tailrec) solve_linear_equation :: "bool \<Rightarrow> real \<Rightarrow> bool list"
   133   where
   134 "solve_linear_equation e_e v_v = (
   135   let
   136     e_e = (
   137       (Try (Rewrite ''all_left'')) #>
   138       (Try (Repeat (Rewrite ''makex1_x''))) #>
   139       (Try (Rewrite_Set ''expand_binoms'')) #>
   140       (Try (Repeat (Rewrite_Set_Inst [(''bdv'', v_v)] ''make_ratpoly_in''))) #>
   141       (Try (Repeat (Rewrite_Set ''LinPoly_simplify''))) ) e_e;
   142     e_e = (
   143       (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''LinEq_simplify'')) #>
   144       (Repeat (Try (Rewrite_Set ''LinPoly_simplify''))) ) e_e
   145   in
   146     Or_to_List e_e)"
   147 
   148 method met_eq_lin : "LinEq/solve_lineq_equation" =
   149   \<open>{rew_ord = "termlessI", rls' = LinEq_erls, prog_rls = Rule_Set.empty, where_rls = LinEq_prls, calc = [],
   150     errpats = [], rew_rls = norm_Poly}\<close>
   151   Program: solve_linear_equation.simps
   152   Given: "equality e_e" "solveFor v_v"
   153   Where: "Not ((lhs e_e) is_polyrat_in v_v)" "((lhs e_e) has_degree_in v_v) = 1"
   154   Find: "solutions v_v'i'"
   155 
   156 section \<open>For trials\<close>
   157 ML \<open>
   158 \<close> ML \<open>
   159 \<close> ML \<open>
   160 \<close>
   161 
   162 end
   163