src/Tools/isac/Knowledge/LinEq.thy
author Walther Neuper <walther.neuper@jku.at>
Tue, 12 May 2020 17:42:29 +0200
changeset 59970 ab1c25c0339a
parent 59903 5037ca1b112b
child 59973 8a46c2e7c27a
permissions -rw-r--r--
shift code from struct.Specify to appropriate locations
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
neuper@52148
    16
  makex1_x:          "a^^^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@37972
    26
val thy = @{theory};
neuper@37972
    27
neuper@37950
    28
val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
walther@59852
    29
  Rule_Set.append_rules "LinEq_prls" Rule_Set.empty 
walther@59878
    30
	     [Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
walther@59878
    31
	      Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches ""),
walther@59878
    32
	      Rule.Eval ("Prog_Expr.lhs"    , Prog_Expr.eval_lhs ""),
walther@59878
    33
	      Rule.Eval ("Prog_Expr.rhs"    , Prog_Expr.eval_rhs ""),
walther@59878
    34
	      Rule.Eval ("Poly.has'_degree'_in", eval_has_degree_in ""),
walther@59878
    35
 	      Rule.Eval ("Poly.is'_polyrat'_in", eval_is_polyrat_in ""),
walther@59878
    36
	      Rule.Eval ("Prog_Expr.occurs'_in", Prog_Expr.eval_occurs_in ""),    
walther@59878
    37
	      Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),
walther@59871
    38
	      Rule.Thm ("not_true",ThmC.numerals_to_Free @{thm not_true}),
walther@59871
    39
	      Rule.Thm ("not_false",ThmC.numerals_to_Free @{thm not_false}),
walther@59871
    40
	      Rule.Thm ("and_true",ThmC.numerals_to_Free @{thm and_true}),
walther@59871
    41
	      Rule.Thm ("and_false",ThmC.numerals_to_Free @{thm and_false}),
walther@59871
    42
	      Rule.Thm ("or_true",ThmC.numerals_to_Free @{thm or_true}),
walther@59871
    43
	      Rule.Thm ("or_false",ThmC.numerals_to_Free @{thm or_false})
neuper@37950
    44
              ];
neuper@37950
    45
(* ----- erls ----- *)
neuper@37950
    46
val LinEq_crls = 
walther@59852
    47
   Rule_Set.append_rules "LinEq_crls" poly_crls
walther@59871
    48
   [Rule.Thm  ("real_assoc_1",ThmC.numerals_to_Free @{thm real_assoc_1})
neuper@37950
    49
    (*		
neuper@37950
    50
     Don't use
walther@59878
    51
     Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),
walther@59878
    52
     Rule.Eval ("Prog_Expr.pow" , (**)eval_binop "#power_"),
neuper@37950
    53
     *)
neuper@37950
    54
    ];
neuper@37950
    55
neuper@37950
    56
(* ----- crls ----- *)
neuper@37950
    57
val LinEq_erls = 
walther@59852
    58
   Rule_Set.append_rules "LinEq_erls" Poly_erls
walther@59871
    59
   [Rule.Thm  ("real_assoc_1",ThmC.numerals_to_Free @{thm real_assoc_1})
neuper@37950
    60
    (*		
neuper@37950
    61
     Don't use
walther@59878
    62
     Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),
walther@59878
    63
     Rule.Eval ("Prog_Expr.pow" , (**)eval_binop "#power_"),
neuper@37950
    64
     *)
neuper@37950
    65
    ];
wneuper@59472
    66
\<close>
wneuper@59472
    67
setup \<open>KEStore_Elems.add_rlss 
wneuper@59472
    68
  [("LinEq_erls", (Context.theory_name @{theory}, LinEq_erls))]\<close>
wneuper@59472
    69
ML \<open>
neuper@37950
    70
    
s1210629013@55444
    71
val LinPoly_simplify = prep_rls'(
walther@59851
    72
  Rule_Def.Repeat {id = "LinPoly_simplify", preconds = [], 
neuper@37950
    73
       rew_ord = ("termlessI",termlessI), 
neuper@37950
    74
       erls = LinEq_erls, 
walther@59851
    75
       srls = Rule_Set.Empty, 
neuper@42451
    76
       calc = [], errpatts = [],
neuper@37950
    77
       rules = [
walther@59871
    78
		Rule.Thm  ("real_assoc_1",ThmC.numerals_to_Free @{thm real_assoc_1}),
walther@59878
    79
		Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_"),
walther@59878
    80
		Rule.Eval ("Groups.minus_class.minus", (**)eval_binop "#sub_"),
walther@59878
    81
		Rule.Eval ("Groups.times_class.times", (**)eval_binop "#mult_"),
neuper@37950
    82
		(*  Dont use  
walther@59878
    83
		 Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),		
walther@59878
    84
		 Rule.Eval ("NthRoot.sqrt", eval_sqrt "#sqrt_"),
neuper@37950
    85
		 *)
walther@59878
    86
		Rule.Eval ("Prog_Expr.pow" , (**)eval_binop "#power_")
neuper@37950
    87
		],
walther@59878
    88
       scr = Rule.Empty_Prog});
wneuper@59472
    89
\<close>
wneuper@59472
    90
setup \<open>KEStore_Elems.add_rlss 
wneuper@59472
    91
  [("LinPoly_simplify", (Context.theory_name @{theory}, LinPoly_simplify))]\<close>
wneuper@59472
    92
ML \<open>
neuper@37950
    93
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@59851
    96
Rule_Def.Repeat {id = "LinEq_simplify", preconds = [],
walther@59857
    97
     rew_ord = ("e_rew_ord", Rewrite_Ord.e_rew_ord),
neuper@37950
    98
     erls = LinEq_erls,
walther@59851
    99
     srls = Rule_Set.Empty,
neuper@42451
   100
     calc = [], errpatts = [],
neuper@37950
   101
     rules = [
walther@59871
   102
	      Rule.Thm("lin_isolate_add1",ThmC.numerals_to_Free @{thm lin_isolate_add1}), 
neuper@37950
   103
	      (* a+bx=0 -> bx=-a *)
walther@59871
   104
	      Rule.Thm("lin_isolate_add2",ThmC.numerals_to_Free @{thm lin_isolate_add2}), 
neuper@37950
   105
	      (* a+ x=0 ->  x=-a *)
walther@59871
   106
	      Rule.Thm("lin_isolate_div",ThmC.numerals_to_Free @{thm lin_isolate_div})    
neuper@37950
   107
	      (*   bx=c -> x=c/b *)  
neuper@37950
   108
	      ],
walther@59878
   109
     scr = Rule.Empty_Prog});
wneuper@59472
   110
\<close>
wneuper@59472
   111
setup \<open>KEStore_Elems.add_rlss 
wneuper@59472
   112
  [("LinEq_simplify", (Context.theory_name @{theory}, LinEq_simplify))]\<close>
neuper@37950
   113
neuper@37950
   114
(*----------------------------- problem types --------------------------------*)
neuper@37950
   115
(* ---------linear----------- *)
wneuper@59472
   116
setup \<open>KEStore_Elems.add_pbts
walther@59903
   117
  [(Specify.prep_pbt thy "pbl_equ_univ_lin" [] Problem.id_empty
s1210629013@55339
   118
      (["LINEAR", "univariate", "equation"],
s1210629013@55339
   119
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55339
   120
          ("#Where" ,["HOL.False", (*WN0509 just detected: this pbl can never be used?!?*)
s1210629013@55339
   121
              "Not( (lhs e_e) is_polyrat_in v_v)",
s1210629013@55339
   122
              "Not( (rhs e_e) is_polyrat_in v_v)",
s1210629013@55339
   123
              "((lhs e_e) has_degree_in v_v)=1",
s1210629013@55339
   124
	            "((rhs e_e) has_degree_in v_v)=1"]),
s1210629013@55339
   125
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59472
   126
        LinEq_prls, SOME "solve (e_e::bool, v_v)", [["LinEq", "solve_lineq_equation"]]))]\<close>
neuper@37950
   127
s1210629013@55373
   128
(*-------------- methods------------------------------------------------------*)
wneuper@59472
   129
setup \<open>KEStore_Elems.add_mets
walther@59903
   130
    [Specify.prep_met thy "met_eqlin" [] Method.id_empty
s1210629013@55373
   131
      (["LinEq"], [],
walther@59852
   132
        {rew_ord' = "tless_true",rls' = Atools_erls,calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty,
s1210629013@55373
   133
          crls = LinEq_crls, errpats = [], nrls = norm_Poly},
wneuper@59545
   134
        @{thm refl})]
wneuper@59473
   135
\<close>
s1210629013@55373
   136
    (* ansprechen mit ["LinEq","solve_univar_equation"] *)
wneuper@59545
   137
wneuper@59504
   138
partial_function (tailrec) solve_linear_equation :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   139
  where
walther@59635
   140
"solve_linear_equation e_e v_v = (
walther@59635
   141
  let
walther@59635
   142
    e_e = (
walther@59637
   143
      (Try (Rewrite ''all_left'')) #>
walther@59637
   144
      (Try (Repeat (Rewrite ''makex1_x''))) #>
walther@59637
   145
      (Try (Rewrite_Set ''expand_binoms'')) #>
walther@59637
   146
      (Try (Repeat (Rewrite_Set_Inst [(''bdv'', v_v)] ''make_ratpoly_in''))) #>
walther@59635
   147
      (Try (Repeat (Rewrite_Set ''LinPoly_simplify''))) ) e_e;
walther@59635
   148
    e_e = (
walther@59637
   149
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''LinEq_simplify'')) #>
walther@59635
   150
      (Repeat (Try (Rewrite_Set ''LinPoly_simplify''))) ) e_e
walther@59635
   151
  in
walther@59635
   152
    Or_to_List e_e)"
wneuper@59473
   153
setup \<open>KEStore_Elems.add_mets
walther@59903
   154
    [Specify.prep_met thy "met_eq_lin" [] Method.id_empty
s1210629013@55373
   155
      (["LinEq","solve_lineq_equation"],
s1210629013@55373
   156
        [("#Given", ["equality e_e", "solveFor v_v"]),
s1210629013@55373
   157
          ("#Where", ["Not ((lhs e_e) is_polyrat_in v_v)", "((lhs e_e)  has_degree_in v_v) = 1"]),
s1210629013@55373
   158
          ("#Find",  ["solutions v_v'i'"])],
walther@59852
   159
        {rew_ord' = "termlessI", rls' = LinEq_erls, srls = Rule_Set.empty, prls = LinEq_prls, calc = [],
s1210629013@55373
   160
          crls = LinEq_crls, errpats = [], nrls = norm_Poly},
wneuper@59551
   161
        @{thm solve_linear_equation.simps})]
wneuper@59472
   162
\<close>
walther@59970
   163
ML \<open>Method.from_store' @{theory} ["LinEq","solve_lineq_equation"];\<close>
neuper@37950
   164
neuper@37906
   165
end
neuper@37906
   166