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