src/Tools/isac/Knowledge/LinEq.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Fri, 23 Mar 2018 10:14:39 +0100
changeset 59411 3e241a6938ce
parent 59406 509d70b507e5
child 59416 229e5c9cf78b
permissions -rw-r--r--
Celem: Test_Isac partially

"xxxe_rew_ordxxx" has slipped in with last changeset.
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@37906
    12
consts
neuper@37906
    13
   Solve'_lineq'_equation
neuper@37950
    14
             :: "[bool,real, 
neuper@37950
    15
		   bool list] => bool list"
wneuper@59334
    16
               ("((Script Solve'_lineq'_equation (_ _ =))// (_))" 9)
neuper@37906
    17
neuper@52148
    18
axiomatization where
wneuper@59370
    19
(*-- normalise --*)
neuper@37906
    20
  (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*)
neuper@52148
    21
  all_left:          "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)" and
neuper@52148
    22
  makex1_x:          "a^^^1  = a" and
neuper@52148
    23
  real_assoc_1:      "a+(b+c) = a+b+c" and
neuper@52148
    24
  real_assoc_2:      "a*(b*c) = a*b*c" and
neuper@37906
    25
neuper@37906
    26
(*-- solve --*)
neuper@52148
    27
  lin_isolate_add1:  "(a + b*bdv = 0) = (b*bdv = (-1)*a)" and
neuper@52148
    28
  lin_isolate_add2:  "(a +   bdv = 0) = (  bdv = (-1)*a)" and
neuper@37982
    29
  lin_isolate_div:   "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)"
neuper@37950
    30
neuper@37950
    31
ML {*
neuper@37972
    32
val thy = @{theory};
neuper@37972
    33
neuper@37950
    34
val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
wneuper@59406
    35
  Celem.append_rls "LinEq_prls" Celem.e_rls 
wneuper@59406
    36
	     [Celem.Calc ("HOL.eq",eval_equal "#equal_"),
wneuper@59406
    37
	      Celem.Calc ("Tools.matches",eval_matches ""),
wneuper@59406
    38
	      Celem.Calc ("Tools.lhs"    ,eval_lhs ""),
wneuper@59406
    39
	      Celem.Calc ("Tools.rhs"    ,eval_rhs ""),
wneuper@59406
    40
	      Celem.Calc ("Poly.has'_degree'_in",eval_has_degree_in ""),
wneuper@59406
    41
 	      Celem.Calc ("Poly.is'_polyrat'_in",eval_is_polyrat_in ""),
wneuper@59406
    42
	      Celem.Calc ("Atools.occurs'_in",eval_occurs_in ""),    
wneuper@59406
    43
	      Celem.Calc ("Atools.ident",eval_ident "#ident_"),
wneuper@59406
    44
	      Celem.Thm ("not_true",TermC.num_str @{thm not_true}),
wneuper@59406
    45
	      Celem.Thm ("not_false",TermC.num_str @{thm not_false}),
wneuper@59406
    46
	      Celem.Thm ("and_true",TermC.num_str @{thm and_true}),
wneuper@59406
    47
	      Celem.Thm ("and_false",TermC.num_str @{thm and_false}),
wneuper@59406
    48
	      Celem.Thm ("or_true",TermC.num_str @{thm or_true}),
wneuper@59406
    49
	      Celem.Thm ("or_false",TermC.num_str @{thm or_false})
neuper@37950
    50
              ];
neuper@37950
    51
(* ----- erls ----- *)
neuper@37950
    52
val LinEq_crls = 
wneuper@59406
    53
   Celem.append_rls "LinEq_crls" poly_crls
wneuper@59406
    54
   [Celem.Thm  ("real_assoc_1",TermC.num_str @{thm real_assoc_1})
neuper@37950
    55
    (*		
neuper@37950
    56
     Don't use
wneuper@59406
    57
     Celem.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),
wneuper@59406
    58
     Celem.Calc ("Atools.pow" ,eval_binop "#power_"),
neuper@37950
    59
     *)
neuper@37950
    60
    ];
neuper@37950
    61
neuper@37950
    62
(* ----- crls ----- *)
neuper@37950
    63
val LinEq_erls = 
wneuper@59406
    64
   Celem.append_rls "LinEq_erls" Poly_erls
wneuper@59406
    65
   [Celem.Thm  ("real_assoc_1",TermC.num_str @{thm real_assoc_1})
neuper@37950
    66
    (*		
neuper@37950
    67
     Don't use
wneuper@59406
    68
     Celem.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),
wneuper@59406
    69
     Celem.Calc ("Atools.pow" ,eval_binop "#power_"),
neuper@37950
    70
     *)
neuper@37950
    71
    ];
neuper@52125
    72
*}
neuper@52125
    73
setup {* KEStore_Elems.add_rlss 
neuper@52125
    74
  [("LinEq_erls", (Context.theory_name @{theory}, LinEq_erls))] *}
neuper@52125
    75
ML {*
neuper@37950
    76
    
s1210629013@55444
    77
val LinPoly_simplify = prep_rls'(
wneuper@59406
    78
  Celem.Rls {id = "LinPoly_simplify", preconds = [], 
neuper@37950
    79
       rew_ord = ("termlessI",termlessI), 
neuper@37950
    80
       erls = LinEq_erls, 
wneuper@59406
    81
       srls = Celem.Erls, 
neuper@42451
    82
       calc = [], errpatts = [],
neuper@37950
    83
       rules = [
wneuper@59406
    84
		Celem.Thm  ("real_assoc_1",TermC.num_str @{thm real_assoc_1}),
wneuper@59406
    85
		Celem.Calc ("Groups.plus_class.plus",eval_binop "#add_"),
wneuper@59406
    86
		Celem.Calc ("Groups.minus_class.minus",eval_binop "#sub_"),
wneuper@59406
    87
		Celem.Calc ("Groups.times_class.times",eval_binop "#mult_"),
neuper@37950
    88
		(*  Dont use  
wneuper@59406
    89
		 Celem.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),		
wneuper@59406
    90
		 Celem.Calc ("NthRoot.sqrt",eval_sqrt "#sqrt_"),
neuper@37950
    91
		 *)
wneuper@59406
    92
		Celem.Calc ("Atools.pow" ,eval_binop "#power_")
neuper@37950
    93
		],
wneuper@59406
    94
       scr = Celem.EmptyScr});
neuper@52125
    95
*}
neuper@52125
    96
setup {* KEStore_Elems.add_rlss 
neuper@52125
    97
  [("LinPoly_simplify", (Context.theory_name @{theory}, LinPoly_simplify))] *}
neuper@52125
    98
ML {*
neuper@37950
    99
neuper@37950
   100
(*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*)
s1210629013@55444
   101
val LinEq_simplify = prep_rls'(
wneuper@59406
   102
Celem.Rls {id = "LinEq_simplify", preconds = [],
wneuper@59411
   103
     rew_ord = ("e_rew_ord", Celem.e_rew_ord),
neuper@37950
   104
     erls = LinEq_erls,
wneuper@59406
   105
     srls = Celem.Erls,
neuper@42451
   106
     calc = [], errpatts = [],
neuper@37950
   107
     rules = [
wneuper@59406
   108
	      Celem.Thm("lin_isolate_add1",TermC.num_str @{thm lin_isolate_add1}), 
neuper@37950
   109
	      (* a+bx=0 -> bx=-a *)
wneuper@59406
   110
	      Celem.Thm("lin_isolate_add2",TermC.num_str @{thm lin_isolate_add2}), 
neuper@37950
   111
	      (* a+ x=0 ->  x=-a *)
wneuper@59406
   112
	      Celem.Thm("lin_isolate_div",TermC.num_str @{thm lin_isolate_div})    
neuper@37950
   113
	      (*   bx=c -> x=c/b *)  
neuper@37950
   114
	      ],
wneuper@59406
   115
     scr = Celem.EmptyScr});
neuper@52125
   116
*}
neuper@52125
   117
setup {* KEStore_Elems.add_rlss 
neuper@52125
   118
  [("LinEq_simplify", (Context.theory_name @{theory}, LinEq_simplify))] *}
neuper@37950
   119
neuper@37950
   120
(*----------------------------- problem types --------------------------------*)
neuper@37950
   121
(* ---------linear----------- *)
s1210629013@55359
   122
setup {* KEStore_Elems.add_pbts
wneuper@59406
   123
  [(Specify.prep_pbt thy "pbl_equ_univ_lin" [] Celem.e_pblID
s1210629013@55339
   124
      (["LINEAR", "univariate", "equation"],
s1210629013@55339
   125
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55339
   126
          ("#Where" ,["HOL.False", (*WN0509 just detected: this pbl can never be used?!?*)
s1210629013@55339
   127
              "Not( (lhs e_e) is_polyrat_in v_v)",
s1210629013@55339
   128
              "Not( (rhs e_e) is_polyrat_in v_v)",
s1210629013@55339
   129
              "((lhs e_e) has_degree_in v_v)=1",
s1210629013@55339
   130
	            "((rhs e_e) has_degree_in v_v)=1"]),
s1210629013@55339
   131
          ("#Find"  ,["solutions v_v'i'"])],
s1210629013@55339
   132
        LinEq_prls, SOME "solve (e_e::bool, v_v)", [["LinEq", "solve_lineq_equation"]]))] *}
neuper@37950
   133
s1210629013@55373
   134
(*-------------- methods------------------------------------------------------*)
s1210629013@55373
   135
setup {* KEStore_Elems.add_mets
wneuper@59406
   136
  [Specify.prep_met thy "met_eqlin" [] Celem.e_metID
s1210629013@55373
   137
      (["LinEq"], [],
wneuper@59406
   138
        {rew_ord' = "tless_true",rls' = Atools_erls,calc = [], srls = Celem.e_rls, prls = Celem.e_rls,
s1210629013@55373
   139
          crls = LinEq_crls, errpats = [], nrls = norm_Poly},
s1210629013@55373
   140
        "empty_script"),
s1210629013@55373
   141
    (* ansprechen mit ["LinEq","solve_univar_equation"] *)
wneuper@59406
   142
    Specify.prep_met thy "met_eq_lin" [] Celem.e_metID
s1210629013@55373
   143
      (["LinEq","solve_lineq_equation"],
s1210629013@55373
   144
        [("#Given", ["equality e_e", "solveFor v_v"]),
s1210629013@55373
   145
          ("#Where", ["Not ((lhs e_e) is_polyrat_in v_v)", "((lhs e_e)  has_degree_in v_v) = 1"]),
s1210629013@55373
   146
          ("#Find",  ["solutions v_v'i'"])],
wneuper@59406
   147
        {rew_ord' = "termlessI", rls' = LinEq_erls, srls = Celem.e_rls, prls = LinEq_prls, calc = [],
s1210629013@55373
   148
          crls = LinEq_crls, errpats = [], nrls = norm_Poly},
s1210629013@55373
   149
        "Script Solve_lineq_equation (e_e::bool) (v_v::real) =                 " ^
s1210629013@55373
   150
          "(let e_e =((Try         (Rewrite      all_left           False)) @@   " ^ 
s1210629013@55373
   151
          "           (Try (Repeat (Rewrite     makex1_x            False))) @@  " ^ 
s1210629013@55373
   152
          "           (Try         (Rewrite_Set expand_binoms       False)) @@   " ^ 
s1210629013@55373
   153
          "           (Try (Repeat (Rewrite_Set_Inst [(bdv, v_v::real)]          " ^
s1210629013@55373
   154
          "                                 make_ratpoly_in    False)))    @@    " ^
s1210629013@55373
   155
          "           (Try (Repeat (Rewrite_Set LinPoly_simplify      False))))e_e;" ^
s1210629013@55373
   156
          "     e_e = ((Try (Rewrite_Set_Inst [(bdv, v_v::real)]                  " ^
s1210629013@55373
   157
          "                                     LinEq_simplify True)) @@  " ^
s1210629013@55373
   158
          "            (Repeat(Try (Rewrite_Set LinPoly_simplify     False)))) e_e " ^
s1210629013@55373
   159
          " in ((Or_to_List e_e)::bool list))")]
s1210629013@55373
   160
*}
wneuper@59269
   161
ML {* Specify.get_met' @{theory} ["LinEq","solve_lineq_equation"]; *}
neuper@37950
   162
neuper@37906
   163
end
neuper@37906
   164