src/Tools/isac/Knowledge/RatEq.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Tue, 19 Feb 2019 19:35:12 +0100
changeset 59504 546bd1b027cc
parent 59491 516e6cc731ab
child 59505 a1f223658994
permissions -rw-r--r--
[-Test_Isac] funpack: cp program code to partial_function

tests are broken, although no Script has been changed -
reason: partial_function introduces constants.
neuper@37906
     1
(*.(c) by Richard Lang, 2003 .*)
neuper@37906
     2
(* theory collecting all knowledge for RationalEquations
neuper@37906
     3
   created by: rlang 
neuper@37906
     4
         date: 02.08.12
neuper@37906
     5
   changed by: rlang
neuper@37906
     6
   last change by: rlang
neuper@37906
     7
             date: 02.11.28
neuper@37906
     8
*)
neuper@37906
     9
s1210629013@55339
    10
theory RatEq imports Rational Equation begin
neuper@37906
    11
wneuper@59472
    12
text \<open>univariate equations over multivariate rational terms:
neuper@42398
    13
  In 2003 this type has been integrated into ISAC's equation solver
neuper@42398
    14
  by Richard Lang; the root for the solver is Equation.thy.
neuper@42398
    15
  The migration Isabelle2002 --> 2011 found that application of theorems like 
neuper@42398
    16
  rat_mult_denominator_right: "[|Not(d=0)|] ==> ((a::real) = c / d) = (a*d = c)"
neuper@42398
    17
  in rule-sets does not transfer "d ~= 0" to the assumptions; see
neuper@42398
    18
  test --- repair NO asms from rls RatEq_eliminate ---.
neuper@42398
    19
  Thus the migration dropped update of Check_elementwise, which would require
neuper@42398
    20
  these assumptions; see 
neuper@42398
    21
  test --- pbl: rational, univariate, equation ---, --- x / (x ^ 2 - 6 * x + 9) - 1 /...
wneuper@59472
    22
\<close>
neuper@42398
    23
neuper@37906
    24
consts
neuper@37906
    25
neuper@37906
    26
  is'_ratequation'_in :: "[bool, real] => bool" ("_ is'_ratequation'_in _")
neuper@37906
    27
neuper@37906
    28
  (*----------------------scripts-----------------------*)
neuper@37906
    29
  Solve'_rat'_equation
neuper@37954
    30
             :: "[bool,real, 
neuper@37954
    31
		   bool list] => bool list"
wneuper@59334
    32
               ("((Script Solve'_rat'_equation (_ _ =))// (_))" 9)
neuper@37906
    33
neuper@52148
    34
axiomatization where
neuper@37906
    35
   (* FIXME also in Poly.thy def. --> FIXED*)
neuper@37906
    36
   (*real_diff_minus            
neuper@37906
    37
   "a - b = a + (-1) * b"*)
neuper@52148
    38
   real_rat_mult_1:   "a*(b/c) = (a*b)/c" and
neuper@52148
    39
   real_rat_mult_2:   "(a/b)*(c/d) = (a*c)/(b*d)" and
neuper@52148
    40
   real_rat_mult_3:   "(a/b)*c = (a*c)/b" and
neuper@52148
    41
   real_rat_pow:      "(a/b)^^^2 = a^^^2/b^^^2" and
neuper@37906
    42
neuper@52148
    43
   rat_double_rat_1:   "[|Not(c=0); Not(d=0)|] ==> (a / (c/d) = (a*d) / c)" and
neuper@37983
    44
   rat_double_rat_2:   "[|Not(b=0);Not(c=0); Not(d=0)|] ==> 
neuper@52148
    45
                                           ((a/b) / (c/d) = (a*d) / (b*c))" and
neuper@52148
    46
   rat_double_rat_3:   "[|Not(b=0);Not(c=0)|] ==> ((a/b) / c = a / (b*c))" and
neuper@37906
    47
neuper@37906
    48
  (* equation to same denominator *)
neuper@37983
    49
  rat_mult_denominator_both:
neuper@52148
    50
   "[|Not(b=0); Not(d=0)|] ==> ((a::real) / b = c / d) = (a*d = c*b)" and
neuper@37983
    51
  rat_mult_denominator_left:
neuper@52148
    52
   "[|Not(d=0)|] ==> ((a::real) = c / d) = (a*d = c)" and
neuper@37983
    53
  rat_mult_denominator_right:
neuper@37906
    54
   "[|Not(b=0)|] ==> ((a::real) / b = c) = (a = c*b)"
neuper@37906
    55
wneuper@59472
    56
ML \<open>
neuper@37972
    57
val thy = @{theory};
neuper@37972
    58
neuper@37954
    59
(*-------------------------functions-----------------------*)
neuper@37954
    60
(* is_rateqation_in becomes true, if a bdv is in the denominator of a fraction*)
neuper@37954
    61
fun is_rateqation_in t v = 
neuper@37954
    62
    let 
wneuper@59389
    63
	fun coeff_in c v = member op = (TermC.vars c) v;
neuper@38031
    64
   	fun finddivide (_ $ _ $ _ $ _) v = error("is_rateqation_in:")
neuper@37954
    65
	    (* at the moment there is no term like this, but ....*)
wneuper@59360
    66
	  | finddivide (t as (Const ("Rings.divide_class.divide",_) $ _ $ b)) v = coeff_in b v
neuper@37954
    67
	  | finddivide (_ $ t1 $ t2) v = (finddivide t1 v) 
neuper@37954
    68
                                         orelse (finddivide t2 v)
neuper@37954
    69
	  | finddivide (_ $ t1) v = (finddivide t1 v)
neuper@37954
    70
	  | finddivide _ _ = false;
neuper@37954
    71
     in
neuper@37954
    72
	finddivide t v
neuper@37954
    73
    end;
neuper@37954
    74
    
neuper@37954
    75
fun eval_is_ratequation_in _ _ 
neuper@37954
    76
       (p as (Const ("RatEq.is'_ratequation'_in",_) $ t $ v)) _  =
neuper@37954
    77
    if is_rateqation_in t v then 
wneuper@59416
    78
	SOME ((Rule.term2str p) ^ " = True",
wneuper@59390
    79
	      HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
wneuper@59416
    80
    else SOME ((Rule.term2str p) ^ " = True",
wneuper@59390
    81
	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@38015
    82
  | eval_is_ratequation_in _ _ _ _ = ((*tracing"### nichts matcht";*) NONE);
neuper@37954
    83
neuper@37954
    84
(*-------------------------rulse-----------------------*)
neuper@37954
    85
val RatEq_prls = (*15.10.02:just the following order due to subterm evaluation*)
wneuper@59416
    86
  Rule.append_rls "RatEq_prls" Rule.e_rls 
wneuper@59416
    87
	     [Rule.Calc ("Atools.ident",eval_ident "#ident_"),
wneuper@59491
    88
	      Rule.Calc ("Tools.matches", Tools.eval_matches ""),
wneuper@59491
    89
	      Rule.Calc ("Tools.lhs"    , Tools.eval_lhs ""),
wneuper@59491
    90
	      Rule.Calc ("Tools.rhs"    , Tools.eval_rhs ""),
wneuper@59416
    91
	      Rule.Calc ("RatEq.is'_ratequation'_in",eval_is_ratequation_in ""),
wneuper@59416
    92
	      Rule.Calc ("HOL.eq",eval_equal "#equal_"),
wneuper@59416
    93
	      Rule.Thm ("not_true",TermC.num_str @{thm not_true}),
wneuper@59416
    94
	      Rule.Thm ("not_false",TermC.num_str @{thm not_false}),
wneuper@59416
    95
	      Rule.Thm ("and_true",TermC.num_str @{thm and_true}),
wneuper@59416
    96
	      Rule.Thm ("and_false",TermC.num_str @{thm and_false}),
wneuper@59416
    97
	      Rule.Thm ("or_true",TermC.num_str @{thm or_true}),
wneuper@59416
    98
	      Rule.Thm ("or_false",TermC.num_str @{thm or_false})
neuper@37954
    99
	      ];
neuper@37954
   100
neuper@37954
   101
wneuper@59416
   102
(*rls = Rule.merge_rls erls Poly_erls *)
neuper@37954
   103
val rateq_erls = 
wneuper@59416
   104
    Rule.remove_rls "rateq_erls"                             (*WN: ein Hack*)
wneuper@59416
   105
	(Rule.merge_rls "is_ratequation_in" calculate_Rational
wneuper@59416
   106
		   (Rule.append_rls "is_ratequation_in"
neuper@37954
   107
			Poly_erls
wneuper@59416
   108
			[(*Rule.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),*)
wneuper@59416
   109
			 Rule.Calc ("RatEq.is'_ratequation'_in",
neuper@37954
   110
			       eval_is_ratequation_in "")
neuper@37954
   111
neuper@37954
   112
			 ]))
wneuper@59416
   113
	[Rule.Thm ("and_commute",TermC.num_str @{thm and_commute}), (*WN: ein Hack*)
wneuper@59416
   114
	 Rule.Thm ("or_commute",TermC.num_str @{thm or_commute})    (*WN: ein Hack*)
neuper@37954
   115
	 ];
wneuper@59472
   116
\<close>
wneuper@59472
   117
setup \<open>KEStore_Elems.add_rlss [("rateq_erls", (Context.theory_name @{theory}, rateq_erls))]\<close>
wneuper@59472
   118
ML \<open>
neuper@37954
   119
neuper@37954
   120
val RatEq_crls = 
wneuper@59416
   121
    Rule.remove_rls "RatEq_crls"                              (*WN: ein Hack*)
wneuper@59416
   122
	(Rule.merge_rls "is_ratequation_in" calculate_Rational
wneuper@59416
   123
		   (Rule.append_rls "is_ratequation_in"
neuper@37954
   124
			Poly_erls
wneuper@59416
   125
			[(*Rule.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),*)
wneuper@59416
   126
			 Rule.Calc ("RatEq.is'_ratequation'_in",
neuper@37954
   127
			       eval_is_ratequation_in "")
neuper@37954
   128
			 ]))
wneuper@59416
   129
	[Rule.Thm ("and_commute",TermC.num_str @{thm and_commute}), (*WN: ein Hack*)
wneuper@59416
   130
	 Rule.Thm ("or_commute",TermC.num_str @{thm or_commute})    (*WN: ein Hack*)
neuper@37954
   131
	 ];
neuper@37954
   132
s1210629013@55444
   133
val RatEq_eliminate = prep_rls'(
wneuper@59416
   134
  Rule.Rls {id = "RatEq_eliminate", preconds = [],
wneuper@59416
   135
       rew_ord = ("termlessI", termlessI), erls = rateq_erls, srls = Rule.Erls, 
neuper@42451
   136
       calc = [], errpatts = [],
neuper@37954
   137
       rules = [
wneuper@59416
   138
	    Rule.Thm("rat_mult_denominator_both",TermC.num_str @{thm rat_mult_denominator_both}), 
neuper@37954
   139
	     (* a/b=c/d -> ad=cb *)
wneuper@59416
   140
	    Rule.Thm("rat_mult_denominator_left",TermC.num_str @{thm rat_mult_denominator_left}), 
neuper@37954
   141
	     (* a  =c/d -> ad=c  *)
wneuper@59416
   142
	    Rule.Thm("rat_mult_denominator_right",TermC.num_str @{thm rat_mult_denominator_right})
neuper@37954
   143
	     (* a/b=c   ->  a=cb *)
neuper@37954
   144
	    ],
wneuper@59416
   145
    scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
wneuper@59406
   146
    });
wneuper@59472
   147
\<close>
wneuper@59472
   148
setup \<open>KEStore_Elems.add_rlss [("RatEq_eliminate",
wneuper@59472
   149
  (Context.theory_name @{theory}, RatEq_eliminate))]\<close>
wneuper@59472
   150
ML \<open>
neuper@37954
   151
s1210629013@55444
   152
val RatEq_simplify = prep_rls'(
wneuper@59416
   153
  Rule.Rls {id = "RatEq_simplify", preconds = [], rew_ord = ("termlessI", termlessI),
wneuper@59416
   154
      erls = rateq_erls, srls = Rule.Erls, calc = [], errpatts = [],
neuper@37954
   155
    rules = [
wneuper@59416
   156
	     Rule.Thm("real_rat_mult_1",TermC.num_str @{thm real_rat_mult_1}),
neuper@37954
   157
	     (*a*(b/c) = (a*b)/c*)
wneuper@59416
   158
	     Rule.Thm("real_rat_mult_2",TermC.num_str @{thm real_rat_mult_2}),
neuper@37954
   159
	     (*(a/b)*(c/d) = (a*c)/(b*d)*)
wneuper@59416
   160
             Rule.Thm("real_rat_mult_3",TermC.num_str @{thm real_rat_mult_3}),
neuper@37954
   161
             (* (a/b)*c = (a*c)/b*)
wneuper@59416
   162
	     Rule.Thm("real_rat_pow",TermC.num_str @{thm real_rat_pow}),
neuper@37954
   163
	     (*(a/b)^^^2 = a^^^2/b^^^2*)
wneuper@59416
   164
	     Rule.Thm("real_diff_minus",TermC.num_str @{thm real_diff_minus}),
neuper@37954
   165
	     (* a - b = a + (-1) * b *)
wneuper@59416
   166
             Rule.Thm("rat_double_rat_1",TermC.num_str @{thm rat_double_rat_1}),
neuper@37954
   167
             (* (a / (c/d) = (a*d) / c) *)
wneuper@59416
   168
             Rule.Thm("rat_double_rat_2",TermC.num_str @{thm rat_double_rat_2}), 
neuper@37954
   169
             (* ((a/b) / (c/d) = (a*d) / (b*c)) *)
wneuper@59416
   170
             Rule.Thm("rat_double_rat_3",TermC.num_str @{thm rat_double_rat_3}) 
neuper@37954
   171
             (* ((a/b) / c = a / (b*c) ) *)
neuper@37954
   172
	     ],
wneuper@59416
   173
    scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
wneuper@59406
   174
    });
wneuper@59472
   175
\<close>
wneuper@59472
   176
setup \<open>KEStore_Elems.add_rlss [("RatEq_simplify",
wneuper@59472
   177
  (Context.theory_name @{theory}, RatEq_simplify))]\<close>
wneuper@59472
   178
ML \<open>
neuper@37954
   179
(*-------------------------Problem-----------------------*)
neuper@37954
   180
(*
neuper@37954
   181
(get_pbt ["rational","univariate","equation"]);
neuper@37954
   182
show_ptyps(); 
neuper@37954
   183
*)
wneuper@59472
   184
\<close>
wneuper@59472
   185
setup \<open>KEStore_Elems.add_pbts
wneuper@59406
   186
  [(Specify.prep_pbt thy "pbl_equ_univ_rat" [] Celem.e_pblID
s1210629013@55339
   187
    (["rational","univariate","equation"],
s1210629013@55339
   188
      [("#Given", ["equality e_e","solveFor v_v"]),
s1210629013@55339
   189
        ("#Where", ["(e_e::bool) is_ratequation_in (v_v::real)"]),
s1210629013@55339
   190
        ("#Find", ["solutions v_v'i'"])],
wneuper@59472
   191
      RatEq_prls, SOME "solve (e_e::bool, v_v)", [["RatEq","solve_rat_equation"]]))]\<close>
neuper@37954
   192
s1210629013@55373
   193
(*-------------------------methods-----------------------*)
wneuper@59472
   194
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   195
    [Specify.prep_met thy "met_rateq" [] Celem.e_metID
s1210629013@55373
   196
      (["RatEq"], [],
wneuper@59416
   197
        {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = Rule.e_rls, prls=Rule.e_rls,
wneuper@59473
   198
          crls=RatEq_crls, errpats = [], nrls = norm_Rational}, "empty_script")]
wneuper@59473
   199
\<close>
wneuper@59504
   200
partial_function (tailrec) solve_rational_equ :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   201
  where
wneuper@59504
   202
"solve_rational_equ e_e v_v =
wneuper@59504
   203
(let e_e = ((Repeat(Try (Rewrite_Set ''RatEq_simplify'' True))) @@
wneuper@59504
   204
            (Repeat(Try (Rewrite_Set ''norm_Rational'' False))) @@
wneuper@59504
   205
            (Repeat(Try (Rewrite_Set ''add_fractions_p'' False))) @@
wneuper@59504
   206
            (Repeat(Try (Rewrite_Set ''RatEq_eliminate'' True)))) e_e;
wneuper@59504
   207
     L_L = SubProblem (''RatEq'', [''univariate'', ''equation''], [''no_met''])
wneuper@59504
   208
            [BOOL e_e, REAL v_v]
wneuper@59504
   209
 in Check_elementwise L_L {(v_v::real). Assumptions})"
wneuper@59473
   210
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   211
    [Specify.prep_met thy "met_rat_eq" [] Celem.e_metID
s1210629013@55373
   212
      (["RatEq", "solve_rat_equation"],
s1210629013@55373
   213
        [("#Given" ,["equality e_e","solveFor v_v"]),
s1210629013@55373
   214
          ("#Where" ,["(e_e::bool) is_ratequation_in (v_v::real)"]),
s1210629013@55373
   215
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59416
   216
        {rew_ord'="termlessI", rls'=rateq_erls, srls=Rule.e_rls, prls=RatEq_prls, calc=[],
s1210629013@55373
   217
          crls=RatEq_crls, errpats = [], nrls = norm_Rational},
s1210629013@55373
   218
        "Script Solve_rat_equation  (e_e::bool) (v_v::real) =                   " ^
wneuper@59489
   219
          "(let e_e = ((Repeat(Try (Rewrite_Set ''RatEq_simplify''      True))) @@  " ^
wneuper@59489
   220
          "           (Repeat(Try (Rewrite_Set ''norm_Rational''      False))) @@  " ^
wneuper@59489
   221
          "           (Repeat(Try (Rewrite_Set ''add_fractions_p'' False))) @@  " ^
wneuper@59489
   222
          "           (Repeat(Try (Rewrite_Set ''RatEq_eliminate''     True)))) e_e;" ^
wneuper@59489
   223
          " (L_L::bool list) = (SubProblem (''RatEq'',[''univariate'',''equation''], [''no_met''])" ^
s1210629013@55373
   224
          "                    [BOOL e_e, REAL v_v])                     " ^
s1210629013@55373
   225
          " in Check_elementwise L_L {(v_v::real). Assumptions})")]
wneuper@59472
   226
\<close>
neuper@37954
   227
wneuper@59472
   228
setup \<open>KEStore_Elems.add_calcs
wneuper@59472
   229
  [("is_ratequation_in", ("RatEq.is_ratequation_in", eval_is_ratequation_in ""))]\<close>
neuper@37906
   230
neuper@37906
   231
end