src/Tools/isac/Knowledge/RatEq.thy
author Walther Neuper <neuper@ist.tugraz.at>
Mon, 16 Sep 2013 12:20:00 +0200
changeset 52105 2786cc9704c8
parent 48789 498ed5bb1004
child 52125 6f1d3415dc68
permissions -rw-r--r--
Test_Isac works again, perfectly ..

# the same tests works as in 8df4b6196660 (the *child* of "Test_Isac works...")
# ..EXCEPT those marked with "exception Div raised"
# for general state of tests see Test_Isac section {* history of tests *}.
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
neuper@37954
    10
theory RatEq imports Rational begin
neuper@37906
    11
neuper@42398
    12
text {* 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 /...
neuper@42398
    22
*}
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"
neuper@37954
    32
               ("((Script Solve'_rat'_equation (_ _ =))// 
neuper@37954
    33
                  (_))" 9)
neuper@37906
    34
t@42211
    35
axioms(*axiomatization where*)
neuper@37906
    36
   (* FIXME also in Poly.thy def. --> FIXED*)
neuper@37906
    37
   (*real_diff_minus            
neuper@37906
    38
   "a - b = a + (-1) * b"*)
t@42211
    39
   real_rat_mult_1:   "a*(b/c) = (a*b)/c" (*and*)
t@42211
    40
   real_rat_mult_2:   "(a/b)*(c/d) = (a*c)/(b*d)" (*and*)
t@42211
    41
   real_rat_mult_3:   "(a/b)*c = (a*c)/b" (*and*)
t@42211
    42
   real_rat_pow:      "(a/b)^^^2 = a^^^2/b^^^2" (*and*)
neuper@37906
    43
t@42211
    44
   rat_double_rat_1:   "[|Not(c=0); Not(d=0)|] ==> (a / (c/d) = (a*d) / c)" (*and*)
neuper@37983
    45
   rat_double_rat_2:   "[|Not(b=0);Not(c=0); Not(d=0)|] ==> 
t@42211
    46
                                           ((a/b) / (c/d) = (a*d) / (b*c))" (*and*)
t@42211
    47
   rat_double_rat_3:   "[|Not(b=0);Not(c=0)|] ==> ((a/b) / c = a / (b*c))" (*and*)
neuper@37906
    48
neuper@37906
    49
  (* equation to same denominator *)
neuper@37983
    50
  rat_mult_denominator_both:
t@42211
    51
   "[|Not(b=0); Not(d=0)|] ==> ((a::real) / b = c / d) = (a*d = c*b)" (*and*)
neuper@37983
    52
  rat_mult_denominator_left:
t@42211
    53
   "[|Not(d=0)|] ==> ((a::real) = c / d) = (a*d = c)" (*and*)
neuper@37983
    54
  rat_mult_denominator_right:
neuper@37906
    55
   "[|Not(b=0)|] ==> ((a::real) / b = c) = (a = c*b)"
neuper@37906
    56
neuper@37954
    57
ML {*
neuper@37972
    58
val thy = @{theory};
neuper@37972
    59
neuper@37954
    60
(*-------------------------functions-----------------------*)
neuper@37954
    61
(* is_rateqation_in becomes true, if a bdv is in the denominator of a fraction*)
neuper@37954
    62
fun is_rateqation_in t v = 
neuper@37954
    63
    let 
neuper@37954
    64
	fun coeff_in c v = member op = (vars c) v;
neuper@38031
    65
   	fun finddivide (_ $ _ $ _ $ _) v = error("is_rateqation_in:")
neuper@37954
    66
	    (* at the moment there is no term like this, but ....*)
neuper@48789
    67
	  | finddivide (t as (Const ("Fields.inverse_class.divide",_) $ _ $ b)) v = coeff_in b v
neuper@37954
    68
	  | finddivide (_ $ t1 $ t2) v = (finddivide t1 v) 
neuper@37954
    69
                                         orelse (finddivide t2 v)
neuper@37954
    70
	  | finddivide (_ $ t1) v = (finddivide t1 v)
neuper@37954
    71
	  | finddivide _ _ = false;
neuper@37954
    72
     in
neuper@37954
    73
	finddivide t v
neuper@37954
    74
    end;
neuper@37954
    75
    
neuper@37954
    76
fun eval_is_ratequation_in _ _ 
neuper@37954
    77
       (p as (Const ("RatEq.is'_ratequation'_in",_) $ t $ v)) _  =
neuper@37954
    78
    if is_rateqation_in t v then 
neuper@37954
    79
	SOME ((term2str p) ^ " = True",
neuper@48760
    80
	      Trueprop $ (mk_equality (p, @{term True})))
neuper@37954
    81
    else SOME ((term2str p) ^ " = True",
neuper@48760
    82
	       Trueprop $ (mk_equality (p, @{term False})))
neuper@38015
    83
  | eval_is_ratequation_in _ _ _ _ = ((*tracing"### nichts matcht";*) NONE);
neuper@37954
    84
neuper@37954
    85
(*-------------------------rulse-----------------------*)
neuper@37954
    86
val RatEq_prls = (*15.10.02:just the following order due to subterm evaluation*)
neuper@37954
    87
  append_rls "RatEq_prls" e_rls 
neuper@37954
    88
	     [Calc ("Atools.ident",eval_ident "#ident_"),
neuper@37954
    89
	      Calc ("Tools.matches",eval_matches ""),
neuper@37954
    90
	      Calc ("Tools.lhs"    ,eval_lhs ""),
neuper@37954
    91
	      Calc ("Tools.rhs"    ,eval_rhs ""),
neuper@37954
    92
	      Calc ("RatEq.is'_ratequation'_in",eval_is_ratequation_in ""),
neuper@41922
    93
	      Calc ("HOL.eq",eval_equal "#equal_"),
neuper@37969
    94
	      Thm ("not_true",num_str @{thm not_true}),
neuper@37969
    95
	      Thm ("not_false",num_str @{thm not_false}),
neuper@37969
    96
	      Thm ("and_true",num_str @{thm and_true}),
neuper@37969
    97
	      Thm ("and_false",num_str @{thm and_false}),
neuper@37969
    98
	      Thm ("or_true",num_str @{thm or_true}),
neuper@37969
    99
	      Thm ("or_false",num_str @{thm or_false})
neuper@37954
   100
	      ];
neuper@37954
   101
neuper@37954
   102
neuper@37954
   103
(*rls = merge_rls erls Poly_erls *)
neuper@37954
   104
val rateq_erls = 
neuper@37954
   105
    remove_rls "rateq_erls"                                   (*WN: ein Hack*)
neuper@37954
   106
	(merge_rls "is_ratequation_in" calculate_Rational
neuper@37954
   107
		   (append_rls "is_ratequation_in"
neuper@37954
   108
			Poly_erls
neuper@48789
   109
			[(*Calc ("Fields.inverse_class.divide", eval_cancel "#divide_e"),*)
neuper@37954
   110
			 Calc ("RatEq.is'_ratequation'_in",
neuper@37954
   111
			       eval_is_ratequation_in "")
neuper@37954
   112
neuper@37954
   113
			 ]))
neuper@37969
   114
	[Thm ("and_commute",num_str @{thm and_commute}), (*WN: ein Hack*)
neuper@37969
   115
	 Thm ("or_commute",num_str @{thm or_commute})    (*WN: ein Hack*)
neuper@37954
   116
	 ];
neuper@37967
   117
ruleset' := overwritelthy @{theory} (!ruleset',
neuper@37954
   118
			[("rateq_erls",rateq_erls)(*FIXXXME:del with rls.rls'*)
neuper@37954
   119
			 ]);
neuper@37954
   120
neuper@37954
   121
neuper@37954
   122
val RatEq_crls = 
neuper@37954
   123
    remove_rls "RatEq_crls"                                   (*WN: ein Hack*)
neuper@37954
   124
	(merge_rls "is_ratequation_in" calculate_Rational
neuper@37954
   125
		   (append_rls "is_ratequation_in"
neuper@37954
   126
			Poly_erls
neuper@48789
   127
			[(*Calc ("Fields.inverse_class.divide", eval_cancel "#divide_e"),*)
neuper@37954
   128
			 Calc ("RatEq.is'_ratequation'_in",
neuper@37954
   129
			       eval_is_ratequation_in "")
neuper@37954
   130
			 ]))
neuper@37969
   131
	[Thm ("and_commute",num_str @{thm and_commute}), (*WN: ein Hack*)
neuper@37969
   132
	 Thm ("or_commute",num_str @{thm or_commute})    (*WN: ein Hack*)
neuper@37954
   133
	 ];
neuper@37954
   134
neuper@37954
   135
val RatEq_eliminate = prep_rls(
neuper@37954
   136
  Rls {id = "RatEq_eliminate", preconds = [],
neuper@37954
   137
       rew_ord = ("termlessI", termlessI), erls = rateq_erls, srls = Erls, 
neuper@42451
   138
       calc = [], errpatts = [],
neuper@37954
   139
       rules = [
neuper@37969
   140
	    Thm("rat_mult_denominator_both",num_str @{thm rat_mult_denominator_both}), 
neuper@37954
   141
	     (* a/b=c/d -> ad=cb *)
neuper@37969
   142
	    Thm("rat_mult_denominator_left",num_str @{thm rat_mult_denominator_left}), 
neuper@37954
   143
	     (* a  =c/d -> ad=c  *)
neuper@37969
   144
	    Thm("rat_mult_denominator_right",num_str @{thm rat_mult_denominator_right})
neuper@37954
   145
	     (* a/b=c   ->  a=cb *)
neuper@37954
   146
	    ],
neuper@48763
   147
    scr = Prog ((term_of o the o (parse thy)) "empty_script")
neuper@37954
   148
    }:rls);
neuper@37967
   149
ruleset' := overwritelthy @{theory} (!ruleset',
neuper@37954
   150
			[("RatEq_eliminate",RatEq_eliminate)
neuper@37954
   151
			 ]);
neuper@37954
   152
neuper@37954
   153
val RatEq_simplify = prep_rls(
neuper@37954
   154
  Rls {id = "RatEq_simplify", preconds = [], rew_ord = ("termlessI", termlessI),
neuper@42451
   155
      erls = rateq_erls, srls = Erls, calc = [], errpatts = [],
neuper@37954
   156
    rules = [
neuper@37969
   157
	     Thm("real_rat_mult_1",num_str @{thm real_rat_mult_1}),
neuper@37954
   158
	     (*a*(b/c) = (a*b)/c*)
neuper@37969
   159
	     Thm("real_rat_mult_2",num_str @{thm real_rat_mult_2}),
neuper@37954
   160
	     (*(a/b)*(c/d) = (a*c)/(b*d)*)
neuper@37969
   161
             Thm("real_rat_mult_3",num_str @{thm real_rat_mult_3}),
neuper@37954
   162
             (* (a/b)*c = (a*c)/b*)
neuper@37969
   163
	     Thm("real_rat_pow",num_str @{thm real_rat_pow}),
neuper@37954
   164
	     (*(a/b)^^^2 = a^^^2/b^^^2*)
neuper@37969
   165
	     Thm("real_diff_minus",num_str @{thm real_diff_minus}),
neuper@37954
   166
	     (* a - b = a + (-1) * b *)
neuper@37969
   167
             Thm("rat_double_rat_1",num_str @{thm rat_double_rat_1}),
neuper@37954
   168
             (* (a / (c/d) = (a*d) / c) *)
neuper@37969
   169
             Thm("rat_double_rat_2",num_str @{thm rat_double_rat_2}), 
neuper@37954
   170
             (* ((a/b) / (c/d) = (a*d) / (b*c)) *)
neuper@37969
   171
             Thm("rat_double_rat_3",num_str @{thm rat_double_rat_3}) 
neuper@37954
   172
             (* ((a/b) / c = a / (b*c) ) *)
neuper@37954
   173
	     ],
neuper@48763
   174
    scr = Prog ((term_of o the o (parse thy)) "empty_script")
neuper@37954
   175
    }:rls);
neuper@37967
   176
ruleset' := overwritelthy @{theory} (!ruleset',
neuper@37954
   177
			[("RatEq_simplify",RatEq_simplify)
neuper@37954
   178
			 ]);
neuper@37954
   179
neuper@37954
   180
(*-------------------------Problem-----------------------*)
neuper@37954
   181
(*
neuper@37954
   182
(get_pbt ["rational","univariate","equation"]);
neuper@37954
   183
show_ptyps(); 
neuper@37954
   184
*)
neuper@37954
   185
store_pbt
neuper@37972
   186
 (prep_pbt thy "pbl_equ_univ_rat" [] e_pblID
neuper@37954
   187
 (["rational","univariate","equation"],
neuper@37981
   188
  [("#Given" ,["equality e_e","solveFor v_v"]),
neuper@37982
   189
   ("#Where" ,["(e_e::bool) is_ratequation_in (v_v::real)"]),
neuper@38012
   190
   ("#Find"  ,["solutions v_v'i'"]) 
neuper@37954
   191
  ],
neuper@37954
   192
neuper@37981
   193
  RatEq_prls, SOME "solve (e_e::bool, v_v)",
neuper@37954
   194
  [["RatEq","solve_rat_equation"]]));
neuper@37988
   195
*}
neuper@37954
   196
neuper@37988
   197
ML {*
neuper@37954
   198
(*-------------------------methods-----------------------*)
neuper@37954
   199
store_met
neuper@37972
   200
 (prep_met thy "met_rateq" [] e_metID
neuper@37954
   201
 (["RatEq"],
neuper@37954
   202
   [],
neuper@37954
   203
   {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = e_rls, prls=e_rls,
neuper@42425
   204
    crls=RatEq_crls, errpats = [], nrls = norm_Rational}, "empty_script"));
neuper@37988
   205
*}
neuper@37988
   206
neuper@37988
   207
ML {*
neuper@37954
   208
store_met
neuper@37972
   209
 (prep_met thy "met_rat_eq" [] e_metID
neuper@42394
   210
 (["RatEq", "solve_rat_equation"],
neuper@37981
   211
   [("#Given" ,["equality e_e","solveFor v_v"]),
neuper@37982
   212
   ("#Where" ,["(e_e::bool) is_ratequation_in (v_v::real)"]),
neuper@38012
   213
   ("#Find"  ,["solutions v_v'i'"])
neuper@37954
   214
  ],
neuper@37954
   215
   {rew_ord'="termlessI",
neuper@37954
   216
    rls'=rateq_erls,
neuper@37954
   217
    srls=e_rls,
neuper@37954
   218
    prls=RatEq_prls,
neuper@37954
   219
    calc=[],
neuper@42425
   220
    crls=RatEq_crls, errpats = [], nrls = norm_Rational},
neuper@37982
   221
   "Script Solve_rat_equation  (e_e::bool) (v_v::real) =                   " ^
neuper@37981
   222
    "(let e_e = ((Repeat(Try (Rewrite_Set RatEq_simplify      True))) @@  " ^
neuper@37954
   223
    "           (Repeat(Try (Rewrite_Set norm_Rational      False))) @@  " ^
neuper@52105
   224
    "           (Repeat(Try (Rewrite_Set add_fractions_p False))) @@  " ^
neuper@37988
   225
    "           (Repeat(Try (Rewrite_Set RatEq_eliminate     True)))) e_e;" ^
neuper@37988
   226
    " (L_L::bool list) = (SubProblem (RatEq',[univariate,equation], [no_met])" ^
neuper@37988
   227
    "                    [BOOL e_e, REAL v_v])                     " ^
neuper@42268
   228
    " in Check_elementwise L_L {(v_v::real). Assumptions})"
neuper@37954
   229
   ));
neuper@37988
   230
*}
neuper@37954
   231
neuper@37988
   232
ML {*
neuper@37954
   233
calclist':= overwritel (!calclist', 
neuper@37954
   234
   [("is_ratequation_in", ("RatEq.is_ratequation_in", 
neuper@37954
   235
			   eval_is_ratequation_in ""))
neuper@37954
   236
    ]);
neuper@37954
   237
*}
neuper@37906
   238
neuper@37906
   239
end