src/Tools/isac/Knowledge/LinEq.thy
author wenzelm
Tue, 15 Jun 2021 22:24:20 +0200
changeset 60303 815b0dc8b589
parent 60297 73e7175a7d3f
child 60306 51ec2e101e9f
permissions -rw-r--r--
Isar command 'method' as combination of KEStore_Elems.add_mets + MethodC.prep_import, without change of semantics;
neuper@37906
     1
(*. (c) by Richard Lang, 2003 .*)
neuper@37906
     2
(* theory collecting all knowledge for LinearEquations
neuper@37906
     3
   created by: rlang 
neuper@37906
     4
         date: 02.10
neuper@37906
     5
   changed by: rlang
neuper@37906
     6
   last change by: rlang
neuper@37906
     7
             date: 02.10.20
neuper@37906
     8
*)
neuper@37906
     9
neuper@37950
    10
theory LinEq imports Poly Equation begin
neuper@37906
    11
neuper@52148
    12
axiomatization where
wneuper@59370
    13
(*-- normalise --*)
neuper@37906
    14
  (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*)
neuper@52148
    15
  all_left:          "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)" and
walther@60242
    16
  makex1_x:          "a \<up> 1  = a" and
neuper@52148
    17
  real_assoc_1:      "a+(b+c) = a+b+c" and
neuper@52148
    18
  real_assoc_2:      "a*(b*c) = a*b*c" and
neuper@37906
    19
neuper@37906
    20
(*-- solve --*)
neuper@52148
    21
  lin_isolate_add1:  "(a + b*bdv = 0) = (b*bdv = (-1)*a)" and
neuper@52148
    22
  lin_isolate_add2:  "(a +   bdv = 0) = (  bdv = (-1)*a)" and
neuper@37982
    23
  lin_isolate_div:   "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)"
neuper@37950
    24
wneuper@59472
    25
ML \<open>
neuper@37950
    26
val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
walther@59852
    27
  Rule_Set.append_rules "LinEq_prls" Rule_Set.empty 
wenzelm@60294
    28
	     [\<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
wenzelm@60294
    29
	      \<^rule_eval>\<open>Prog_Expr.matches\<close> (Prog_Expr.eval_matches ""),
wenzelm@60294
    30
	      \<^rule_eval>\<open>Prog_Expr.lhs\<close> (Prog_Expr.eval_lhs ""),
wenzelm@60294
    31
	      \<^rule_eval>\<open>Prog_Expr.rhs\<close> (Prog_Expr.eval_rhs ""),
wenzelm@60294
    32
	      \<^rule_eval>\<open>has_degree_in\<close> (eval_has_degree_in ""),
wenzelm@60294
    33
 	      \<^rule_eval>\<open>is_polyrat_in\<close> (eval_is_polyrat_in ""),
wenzelm@60294
    34
	      \<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in ""),    
wenzelm@60294
    35
	      \<^rule_eval>\<open>Prog_Expr.ident\<close> (Prog_Expr.eval_ident "#ident_"),
wenzelm@60297
    36
	      \<^rule_thm>\<open>not_true\<close>,
wenzelm@60297
    37
	      \<^rule_thm>\<open>not_false\<close>,
wenzelm@60297
    38
	      \<^rule_thm>\<open>and_true\<close>,
wenzelm@60297
    39
	      \<^rule_thm>\<open>and_false\<close>,
wenzelm@60297
    40
	      \<^rule_thm>\<open>or_true\<close>,
wenzelm@60297
    41
	      \<^rule_thm>\<open>or_false\<close>
neuper@37950
    42
              ];
neuper@37950
    43
(* ----- erls ----- *)
neuper@37950
    44
val LinEq_crls = 
walther@59852
    45
   Rule_Set.append_rules "LinEq_crls" poly_crls
wenzelm@60297
    46
   [\<^rule_thm>\<open>real_assoc_1\<close>
neuper@37950
    47
    (*		
neuper@37950
    48
     Don't use
wenzelm@60294
    49
     \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
wenzelm@60294
    50
     \<^rule_eval>\<open>powr\<close> (**)(eval_binop "#power_"),
neuper@37950
    51
     *)
neuper@37950
    52
    ];
neuper@37950
    53
neuper@37950
    54
(* ----- crls ----- *)
neuper@37950
    55
val LinEq_erls = 
walther@59852
    56
   Rule_Set.append_rules "LinEq_erls" Poly_erls
wenzelm@60297
    57
   [\<^rule_thm>\<open>real_assoc_1\<close>
neuper@37950
    58
    (*		
neuper@37950
    59
     Don't use
wenzelm@60294
    60
     \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
wenzelm@60294
    61
     \<^rule_eval>\<open>powr\<close> (**)(eval_binop "#power_"),
neuper@37950
    62
     *)
neuper@37950
    63
    ];
wneuper@59472
    64
\<close>
wenzelm@60289
    65
rule_set_knowledge LinEq_erls = LinEq_erls
wneuper@59472
    66
ML \<open>
neuper@37950
    67
    
s1210629013@55444
    68
val LinPoly_simplify = prep_rls'(
walther@59851
    69
  Rule_Def.Repeat {id = "LinPoly_simplify", preconds = [], 
neuper@37950
    70
       rew_ord = ("termlessI",termlessI), 
neuper@37950
    71
       erls = LinEq_erls, 
walther@59851
    72
       srls = Rule_Set.Empty, 
neuper@42451
    73
       calc = [], errpatts = [],
neuper@37950
    74
       rules = [
wenzelm@60297
    75
		\<^rule_thm>\<open>real_assoc_1\<close>,
wenzelm@60294
    76
		\<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
wenzelm@60294
    77
		\<^rule_eval>\<open>minus\<close> (**)(eval_binop "#sub_"),
wenzelm@60294
    78
		\<^rule_eval>\<open>times\<close> (**)(eval_binop "#mult_"),
wenzelm@60294
    79
		(*  Don't use  
wenzelm@60294
    80
		 \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
wenzelm@60294
    81
		 \<^rule_eval>\<open>sqrt\<close> (eval_sqrt "#sqrt_"),
neuper@37950
    82
		 *)
wenzelm@60294
    83
		\<^rule_eval>\<open>powr\<close> (**)(eval_binop "#power_")
neuper@37950
    84
		],
walther@59878
    85
       scr = Rule.Empty_Prog});
wneuper@59472
    86
\<close>
wenzelm@60289
    87
rule_set_knowledge LinPoly_simplify = LinPoly_simplify
wneuper@59472
    88
ML \<open>
neuper@37950
    89
neuper@37950
    90
(*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*)
s1210629013@55444
    91
val LinEq_simplify = prep_rls'(
walther@59851
    92
Rule_Def.Repeat {id = "LinEq_simplify", preconds = [],
walther@59857
    93
     rew_ord = ("e_rew_ord", Rewrite_Ord.e_rew_ord),
neuper@37950
    94
     erls = LinEq_erls,
walther@59851
    95
     srls = Rule_Set.Empty,
neuper@42451
    96
     calc = [], errpatts = [],
neuper@37950
    97
     rules = [
wenzelm@60297
    98
	      \<^rule_thm>\<open>lin_isolate_add1\<close>, 
neuper@37950
    99
	      (* a+bx=0 -> bx=-a *)
wenzelm@60297
   100
	      \<^rule_thm>\<open>lin_isolate_add2\<close>, 
neuper@37950
   101
	      (* a+ x=0 ->  x=-a *)
wenzelm@60297
   102
	      \<^rule_thm>\<open>lin_isolate_div\<close>    
neuper@37950
   103
	      (*   bx=c -> x=c/b *)  
neuper@37950
   104
	      ],
walther@59878
   105
     scr = Rule.Empty_Prog});
wneuper@59472
   106
\<close>
wenzelm@60289
   107
rule_set_knowledge LinEq_simplify = LinEq_simplify
neuper@37950
   108
neuper@37950
   109
(*----------------------------- problem types --------------------------------*)
neuper@37950
   110
(* ---------linear----------- *)
wneuper@59472
   111
setup \<open>KEStore_Elems.add_pbts
wenzelm@60290
   112
  [(Problem.prep_input @{theory} "pbl_equ_univ_lin" [] Problem.id_empty
s1210629013@55339
   113
      (["LINEAR", "univariate", "equation"],
s1210629013@55339
   114
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55339
   115
          ("#Where" ,["HOL.False", (*WN0509 just detected: this pbl can never be used?!?*)
s1210629013@55339
   116
              "Not( (lhs e_e) is_polyrat_in v_v)",
s1210629013@55339
   117
              "Not( (rhs e_e) is_polyrat_in v_v)",
s1210629013@55339
   118
              "((lhs e_e) has_degree_in v_v)=1",
s1210629013@55339
   119
	            "((rhs e_e) has_degree_in v_v)=1"]),
s1210629013@55339
   120
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59472
   121
        LinEq_prls, SOME "solve (e_e::bool, v_v)", [["LinEq", "solve_lineq_equation"]]))]\<close>
neuper@37950
   122
s1210629013@55373
   123
(*-------------- methods------------------------------------------------------*)
wenzelm@60303
   124
method met_eqlin : "LinEq" =
wenzelm@60303
   125
  \<open>{rew_ord' = "tless_true",rls' = Atools_erls,calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty,
wenzelm@60303
   126
    crls = LinEq_crls, errpats = [], nrls = norm_Poly}\<close>
walther@59997
   127
    (* ansprechen mit ["LinEq", "solve_univar_equation"] *)
wneuper@59545
   128
wneuper@59504
   129
partial_function (tailrec) solve_linear_equation :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   130
  where
walther@59635
   131
"solve_linear_equation e_e v_v = (
walther@59635
   132
  let
walther@59635
   133
    e_e = (
walther@59637
   134
      (Try (Rewrite ''all_left'')) #>
walther@59637
   135
      (Try (Repeat (Rewrite ''makex1_x''))) #>
walther@59637
   136
      (Try (Rewrite_Set ''expand_binoms'')) #>
walther@59637
   137
      (Try (Repeat (Rewrite_Set_Inst [(''bdv'', v_v)] ''make_ratpoly_in''))) #>
walther@59635
   138
      (Try (Repeat (Rewrite_Set ''LinPoly_simplify''))) ) e_e;
walther@59635
   139
    e_e = (
walther@59637
   140
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''LinEq_simplify'')) #>
walther@59635
   141
      (Repeat (Try (Rewrite_Set ''LinPoly_simplify''))) ) e_e
walther@59635
   142
  in
walther@59635
   143
    Or_to_List e_e)"
wenzelm@60303
   144
wenzelm@60303
   145
method met_eq_lin : "LinEq/solve_lineq_equation" =
wenzelm@60303
   146
  \<open>{rew_ord' = "termlessI", rls' = LinEq_erls, srls = Rule_Set.empty, prls = LinEq_prls, calc = [],
wenzelm@60303
   147
    crls = LinEq_crls, errpats = [], nrls = norm_Poly}\<close>
wenzelm@60303
   148
  Program: solve_linear_equation.simps
wenzelm@60303
   149
  Given: "equality e_e" "solveFor v_v"
wenzelm@60303
   150
  Where: "Not ((lhs e_e) is_polyrat_in v_v)" "((lhs e_e)  has_degree_in v_v) = 1"
wenzelm@60303
   151
  Find: "solutions v_v'i'"
wenzelm@60303
   152
walther@60278
   153
ML \<open>
walther@60278
   154
  MethodC.from_store' @{theory} ["LinEq", "solve_lineq_equation"];
walther@60278
   155
\<close> ML \<open>
walther@60278
   156
\<close> ML \<open>
walther@60278
   157
\<close>
neuper@37950
   158
neuper@37906
   159
end
neuper@37906
   160