src/Tools/isac/Knowledge/PolyMinus.thy
author Walther Neuper <walther.neuper@jku.at>
Fri, 10 Apr 2020 18:32:36 +0200
changeset 59868 d77aa0992e0f
parent 59861 65ec9f679c3f
child 59871 82428ca0d23e
permissions -rw-r--r--
use "UnparseC" for renaming identifiers
neuper@37906
     1
(* attempts to perserve binary minus as wanted by Austrian teachers
neuper@37906
     2
   WN071207
neuper@37906
     3
   (c) due to copyright terms
neuper@37906
     4
*)
neuper@37906
     5
neuper@37950
     6
theory PolyMinus imports (*Poly// due to "is_ratpolyexp" in...*) Rational begin
neuper@37906
     7
neuper@37906
     8
consts
neuper@37906
     9
neuper@37906
    10
  (*predicates for conditions in rewriting*)
neuper@37906
    11
  kleiner     :: "['a, 'a] => bool" 	("_ kleiner _") 
neuper@37906
    12
  ist'_monom  :: "'a => bool"		("_ ist'_monom")
neuper@37906
    13
neuper@37906
    14
  (*the CAS-command*)
neuper@37906
    15
  Probe       :: "[bool, bool list] => bool"  
neuper@37906
    16
	(*"Probe (3*a+2*b+a = 4*a+2*b) [a=1,b=2]"*)
neuper@37906
    17
neuper@37906
    18
  (*descriptions for the pbl and met*)
neuper@37950
    19
  Pruefe      :: "bool => una"
neuper@37950
    20
  mitWert     :: "bool list => tobooll"
neuper@37950
    21
  Geprueft    :: "bool => una"
neuper@37906
    22
neuper@52148
    23
axiomatization where
neuper@37906
    24
neuper@52148
    25
  null_minus:            "0 - a = -a" and
neuper@52148
    26
  vor_minus_mal:         "- a * b = (-a) * b" and
neuper@37906
    27
neuper@37906
    28
  (*commute with invariant (a.b).c -association*)
neuper@37980
    29
  tausche_plus:		"[| b ist_monom; a kleiner b  |] ==> 
neuper@52148
    30
			 (b + a) = (a + b)" and
neuper@37980
    31
  tausche_minus:		"[| b ist_monom; a kleiner b  |] ==> 
neuper@52148
    32
			 (b - a) = (-a + b)" and
neuper@37980
    33
  tausche_vor_plus:	"[| b ist_monom; a kleiner b  |] ==> 
neuper@52148
    34
			 (- b + a) = (a - b)" and
neuper@37980
    35
  tausche_vor_minus:	"[| b ist_monom; a kleiner b  |] ==> 
neuper@52148
    36
			 (- b - a) = (-a - b)" and
neuper@52148
    37
  tausche_plus_plus:	"b kleiner c ==> (a + c + b) = (a + b + c)" and
neuper@52148
    38
  tausche_plus_minus:	"b kleiner c ==> (a + c - b) = (a - b + c)" and
neuper@52148
    39
  tausche_minus_plus:	"b kleiner c ==> (a - c + b) = (a + b - c)" and
neuper@52148
    40
  tausche_minus_minus:	"b kleiner c ==> (a - c - b) = (a - b - c)" and
neuper@37906
    41
neuper@37906
    42
  (*commute with invariant (a.b).c -association*)
neuper@37980
    43
  tausche_mal:		"[| b is_atom; a kleiner b  |] ==> 
neuper@52148
    44
			 (b * a) = (a * b)" and
neuper@37980
    45
  tausche_vor_mal:	"[| b is_atom; a kleiner b  |] ==> 
neuper@52148
    46
			 (-b * a) = (-a * b)" and
neuper@37980
    47
  tausche_mal_mal:	"[| c is_atom; b kleiner c  |] ==> 
neuper@52148
    48
			 (x * c * b) = (x * b * c)" and
neuper@52148
    49
  x_quadrat:             "(x * a) * a = x * a ^^^ 2" and
neuper@37906
    50
neuper@37906
    51
neuper@37980
    52
  subtrahiere:               "[| l is_const; m is_const |] ==>  
neuper@52148
    53
			     m * v - l * v = (m - l) * v" and
neuper@37980
    54
  subtrahiere_von_1:         "[| l is_const |] ==>  
neuper@52148
    55
			     v - l * v = (1 - l) * v" and
neuper@37980
    56
  subtrahiere_1:             "[| l is_const; m is_const |] ==>  
neuper@52148
    57
			     m * v - v = (m - 1) * v" and
neuper@37906
    58
neuper@37980
    59
  subtrahiere_x_plus_minus:  "[| l is_const; m is_const |] ==>  
neuper@52148
    60
			     (x + m * v) - l * v = x + (m - l) * v" and
neuper@37980
    61
  subtrahiere_x_plus1_minus: "[| l is_const |] ==>  
neuper@52148
    62
			     (x + v) - l * v = x + (1 - l) * v" and
neuper@37980
    63
  subtrahiere_x_plus_minus1: "[| m is_const |] ==>  
neuper@52148
    64
			     (x + m * v) - v = x + (m - 1) * v" and
neuper@37906
    65
neuper@37980
    66
  subtrahiere_x_minus_plus:  "[| l is_const; m is_const |] ==>  
neuper@52148
    67
			     (x - m * v) + l * v = x + (-m + l) * v" and
neuper@37980
    68
  subtrahiere_x_minus1_plus: "[| l is_const |] ==>  
neuper@52148
    69
			     (x - v) + l * v = x + (-1 + l) * v" and
neuper@37980
    70
  subtrahiere_x_minus_plus1: "[| m is_const |] ==>  
neuper@52148
    71
			     (x - m * v) + v = x + (-m + 1) * v" and
neuper@37906
    72
neuper@37980
    73
  subtrahiere_x_minus_minus: "[| l is_const; m is_const |] ==>  
neuper@52148
    74
			     (x - m * v) - l * v = x + (-m - l) * v" and
neuper@37980
    75
  subtrahiere_x_minus1_minus:"[| l is_const |] ==>  
neuper@52148
    76
			     (x - v) - l * v = x + (-1 - l) * v" and
neuper@37980
    77
  subtrahiere_x_minus_minus1:"[| m is_const |] ==>  
neuper@52148
    78
			     (x - m * v) - v = x + (-m - 1) * v" and
neuper@37906
    79
neuper@37906
    80
neuper@37980
    81
  addiere_vor_minus:         "[| l is_const; m is_const |] ==>  
neuper@52148
    82
			     - (l * v) +  m * v = (-l + m) * v" and
neuper@37980
    83
  addiere_eins_vor_minus:    "[| m is_const |] ==>  
neuper@52148
    84
			     -  v +  m * v = (-1 + m) * v" and
neuper@37980
    85
  subtrahiere_vor_minus:     "[| l is_const; m is_const |] ==>  
neuper@52148
    86
			     - (l * v) -  m * v = (-l - m) * v" and
neuper@37980
    87
  subtrahiere_eins_vor_minus:"[| m is_const |] ==>  
neuper@52148
    88
			     -  v -  m * v = (-1 - m) * v" and
neuper@37906
    89
neuper@52148
    90
  vorzeichen_minus_weg1:      "l kleiner 0 ==> a + l * b = a - -1*l * b" and
neuper@52148
    91
  vorzeichen_minus_weg2:      "l kleiner 0 ==> a - l * b = a + -1*l * b" and
neuper@52148
    92
  vorzeichen_minus_weg3:      "l kleiner 0 ==> k + a - l * b = k + a + -1*l * b" and
neuper@52148
    93
  vorzeichen_minus_weg4:      "l kleiner 0 ==> k - a - l * b = k - a + -1*l * b" and
neuper@37906
    94
neuper@37974
    95
  (*klammer_plus_plus = (add_assoc RS sym)*)
neuper@52148
    96
  klammer_plus_minus:          "a + (b - c) = (a + b) - c" and
neuper@52148
    97
  klammer_minus_plus:          "a - (b + c) = (a - b) - c" and
neuper@52148
    98
  klammer_minus_minus:         "a - (b - c) = (a - b) + c" and
neuper@37906
    99
neuper@52148
   100
  klammer_mult_minus:          "a * (b - c) = a * b - a * c" and
neuper@37980
   101
  klammer_minus_mult:          "(b - c) * a = b * a - c * a"
neuper@37906
   102
wneuper@59472
   103
ML \<open>
neuper@37972
   104
val thy = @{theory};
neuper@37972
   105
neuper@37950
   106
(** eval functions **)
neuper@37906
   107
neuper@37950
   108
(*. get the identifier from specific monomials; see fun ist_monom .*)
neuper@37950
   109
(*HACK.WN080107*)
neuper@37950
   110
fun increase str = 
neuper@40836
   111
    let val s::ss = Symbol.explode str
neuper@37950
   112
    in implode ((chr (ord s + 1))::ss) end;
neuper@37950
   113
fun identifier (Free (id,_)) = id                            (* 2,   a   *)
neuper@38034
   114
  | identifier (Const ("Groups.times_class.times", _) $ Free (num, _) $ Free (id, _)) = 
neuper@37950
   115
    id                                                       (* 2*a, a*b *)
neuper@38034
   116
  | identifier (Const ("Groups.times_class.times", _) $                          (* 3*a*b    *)
neuper@38034
   117
		     (Const ("Groups.times_class.times", _) $
neuper@37950
   118
			    Free (num, _) $ Free _) $ Free (id, _)) = 
wneuper@59390
   119
    if TermC.is_num' num then id
neuper@37950
   120
    else "|||||||||||||"
walther@59603
   121
  | identifier (Const ("Prog_Expr.pow", _) $ Free (base, _) $ Free (exp, _)) =
wneuper@59390
   122
    if TermC.is_num' base then "|||||||||||||"                  (* a^2      *)
neuper@37950
   123
    else (*increase*) base
neuper@38034
   124
  | identifier (Const ("Groups.times_class.times", _) $ Free (num, _) $          (* 3*a^2    *)
walther@59603
   125
		     (Const ("Prog_Expr.pow", _) $
neuper@37950
   126
			    Free (base, _) $ Free (exp, _))) = 
wneuper@59390
   127
    if TermC.is_num' num andalso not (TermC.is_num' base) then (*increase*) base
neuper@37950
   128
    else "|||||||||||||"
neuper@37950
   129
  | identifier _ = "|||||||||||||"(*the "largest" string*);
neuper@37950
   130
neuper@37950
   131
(*("kleiner", ("PolyMinus.kleiner", eval_kleiner ""))*)
neuper@37950
   132
(* order "by alphabet" w.r.t. var: num < (var | num*var) > (var*var | ..) *)
neuper@37950
   133
fun eval_kleiner _ _ (p as (Const ("PolyMinus.kleiner",_) $ a $ b)) _  =
wneuper@59389
   134
     if TermC.is_num b then
wneuper@59389
   135
	 if TermC.is_num a then (*123 kleiner 32 = True !!!*)
wneuper@59392
   136
	     if TermC.num_of_term a < TermC.num_of_term b then 
walther@59868
   137
		 SOME ((UnparseC.term p) ^ " = True",
wneuper@59390
   138
		       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   139
	     else SOME ((UnparseC.term p) ^ " = False",
wneuper@59390
   140
			HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@37950
   141
	 else (* -1 * -2 kleiner 0 *)
walther@59868
   142
	     SOME ((UnparseC.term p) ^ " = False",
wneuper@59390
   143
		   HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@37950
   144
    else
neuper@37950
   145
	if identifier a < identifier b then 
walther@59868
   146
	     SOME ((UnparseC.term p) ^ " = True",
wneuper@59390
   147
		  HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   148
	else SOME ((UnparseC.term p) ^ " = False",
wneuper@59390
   149
		   HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@37950
   150
  | eval_kleiner _ _ _ _ =  NONE;
neuper@37950
   151
neuper@37950
   152
fun ist_monom (Free (id,_)) = true
neuper@38034
   153
  | ist_monom (Const ("Groups.times_class.times", _) $ Free (num, _) $ Free (id, _)) = 
wneuper@59390
   154
    if TermC.is_num' num then true else false
neuper@37950
   155
  | ist_monom _ = false;
neuper@37950
   156
(*. this function only accepts the most simple monoms       vvvvvvvvvv .*)
neuper@37950
   157
fun ist_monom (Free (id,_)) = true                          (* 2,   a   *)
neuper@38034
   158
  | ist_monom (Const ("Groups.times_class.times", _) $ Free _ $ Free (id, _)) = (* 2*a, a*b *)
wneuper@59390
   159
    if TermC.is_num' id then false else true
neuper@38034
   160
  | ist_monom (Const ("Groups.times_class.times", _) $                          (* 3*a*b    *)
neuper@38034
   161
		     (Const ("Groups.times_class.times", _) $
neuper@37950
   162
			    Free (num, _) $ Free _) $ Free (id, _)) =
wneuper@59390
   163
    if TermC.is_num' num andalso not (TermC.is_num' id) then true else false
walther@59603
   164
  | ist_monom (Const ("Prog_Expr.pow", _) $ Free (base, _) $ Free (exp, _)) = 
neuper@37950
   165
    true                                                    (* a^2      *)
neuper@38034
   166
  | ist_monom (Const ("Groups.times_class.times", _) $ Free (num, _) $          (* 3*a^2    *)
walther@59603
   167
		     (Const ("Prog_Expr.pow", _) $
neuper@37950
   168
			    Free (base, _) $ Free (exp, _))) = 
wneuper@59390
   169
    if TermC.is_num' num then true else false
neuper@37950
   170
  | ist_monom _ = false;
neuper@37950
   171
neuper@37950
   172
(* is this a univariate monomial ? *)
neuper@37950
   173
(*("ist_monom", ("PolyMinus.ist'_monom", eval_ist_monom ""))*)
neuper@37950
   174
fun eval_ist_monom _ _ (p as (Const ("PolyMinus.ist'_monom",_) $ a)) _  =
neuper@37950
   175
    if ist_monom a  then 
walther@59868
   176
	SOME ((UnparseC.term p) ^ " = True",
wneuper@59390
   177
	      HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   178
    else SOME ((UnparseC.term p) ^ " = False",
wneuper@59390
   179
	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@37950
   180
  | eval_ist_monom _ _ _ _ =  NONE;
neuper@37950
   181
neuper@37950
   182
neuper@37950
   183
(** rewrite order **)
neuper@37950
   184
neuper@37950
   185
(** rulesets **)
neuper@37950
   186
neuper@37950
   187
val erls_ordne_alphabetisch =
walther@59852
   188
    Rule_Set.append_rules "erls_ordne_alphabetisch" Rule_Set.empty
walther@59773
   189
	       [Rule.Num_Calc ("PolyMinus.kleiner", eval_kleiner ""),
walther@59773
   190
		Rule.Num_Calc ("PolyMinus.ist'_monom", eval_ist_monom "")
neuper@37950
   191
		];
neuper@37950
   192
neuper@37950
   193
val ordne_alphabetisch = 
walther@59851
   194
  Rule_Def.Repeat{id = "ordne_alphabetisch", preconds = [], 
walther@59857
   195
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), srls = Rule_Set.Empty, calc = [], errpatts = [],
neuper@37950
   196
      erls = erls_ordne_alphabetisch, 
wneuper@59416
   197
      rules = [Rule.Thm ("tausche_plus",TermC.num_str @{thm tausche_plus}),
neuper@37950
   198
	       (*"b kleiner a ==> (b + a) = (a + b)"*)
wneuper@59416
   199
	       Rule.Thm ("tausche_minus",TermC.num_str @{thm tausche_minus}),
neuper@37950
   200
	       (*"b kleiner a ==> (b - a) = (-a + b)"*)
wneuper@59416
   201
	       Rule.Thm ("tausche_vor_plus",TermC.num_str @{thm tausche_vor_plus}),
neuper@37950
   202
	       (*"[| b ist_monom; a kleiner b  |] ==> (- b + a) = (a - b)"*)
wneuper@59416
   203
	       Rule.Thm ("tausche_vor_minus",TermC.num_str @{thm tausche_vor_minus}),
neuper@37950
   204
	       (*"[| b ist_monom; a kleiner b  |] ==> (- b - a) = (-a - b)"*)
wneuper@59416
   205
	       Rule.Thm ("tausche_plus_plus",TermC.num_str @{thm tausche_plus_plus}),
neuper@37950
   206
	       (*"c kleiner b ==> (a + c + b) = (a + b + c)"*)
wneuper@59416
   207
	       Rule.Thm ("tausche_plus_minus",TermC.num_str @{thm tausche_plus_minus}),
neuper@37950
   208
	       (*"c kleiner b ==> (a + c - b) = (a - b + c)"*)
wneuper@59416
   209
	       Rule.Thm ("tausche_minus_plus",TermC.num_str @{thm tausche_minus_plus}),
neuper@37950
   210
	       (*"c kleiner b ==> (a - c + b) = (a + b - c)"*)
wneuper@59416
   211
	       Rule.Thm ("tausche_minus_minus",TermC.num_str @{thm tausche_minus_minus})
neuper@37950
   212
	       (*"c kleiner b ==> (a - c - b) = (a - b - c)"*)
wneuper@59416
   213
	       ], scr = Rule.EmptyScr};
neuper@37950
   214
neuper@37950
   215
val fasse_zusammen = 
walther@59851
   216
    Rule_Def.Repeat{id = "fasse_zusammen", preconds = [], 
walther@59857
   217
	rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
   218
	erls = Rule_Set.append_rules "erls_fasse_zusammen" Rule_Set.empty 
walther@59773
   219
			  [Rule.Num_Calc ("Prog_Expr.is'_const", Prog_Expr.eval_const "#is_const_")], 
walther@59851
   220
	srls = Rule_Set.Empty, calc = [], errpatts = [],
neuper@37950
   221
	rules = 
wneuper@59416
   222
	[Rule.Thm ("real_num_collect",TermC.num_str @{thm real_num_collect}), 
neuper@37950
   223
	 (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
wneuper@59416
   224
	 Rule.Thm ("real_num_collect_assoc_r",TermC.num_str @{thm real_num_collect_assoc_r}),
neuper@37950
   225
	 (*"[| l is_const; m..|] ==>  (k + m * n) + l * n = k + (l + m)*n"*)
wneuper@59416
   226
	 Rule.Thm ("real_one_collect",TermC.num_str @{thm real_one_collect}),	
neuper@37950
   227
	 (*"m is_const ==> n + m * n = (1 + m) * n"*)
wneuper@59416
   228
	 Rule.Thm ("real_one_collect_assoc_r",TermC.num_str @{thm real_one_collect_assoc_r}), 
neuper@37950
   229
	 (*"m is_const ==> (k + n) + m * n = k + (m + 1) * n"*)
neuper@37950
   230
neuper@37950
   231
wneuper@59416
   232
	 Rule.Thm ("subtrahiere",TermC.num_str @{thm subtrahiere}),
neuper@37950
   233
	 (*"[| l is_const; m is_const |] ==> m * v - l * v = (m - l) * v"*)
wneuper@59416
   234
	 Rule.Thm ("subtrahiere_von_1",TermC.num_str @{thm subtrahiere_von_1}),
neuper@37950
   235
	 (*"[| l is_const |] ==> v - l * v = (1 - l) * v"*)
wneuper@59416
   236
	 Rule.Thm ("subtrahiere_1",TermC.num_str @{thm subtrahiere_1}),
neuper@37950
   237
	 (*"[| l is_const; m is_const |] ==> m * v - v = (m - 1) * v"*)
neuper@37950
   238
wneuper@59416
   239
	 Rule.Thm ("subtrahiere_x_plus_minus",TermC.num_str @{thm subtrahiere_x_plus_minus}), 
neuper@37950
   240
	 (*"[| l is_const; m..|] ==> (k + m * n) - l * n = k + ( m - l) * n"*)
wneuper@59416
   241
	 Rule.Thm ("subtrahiere_x_plus1_minus",TermC.num_str @{thm subtrahiere_x_plus1_minus}),
neuper@37950
   242
	 (*"[| l is_const |] ==> (x + v) - l * v = x + (1 - l) * v"*)
wneuper@59416
   243
	 Rule.Thm ("subtrahiere_x_plus_minus1",TermC.num_str @{thm subtrahiere_x_plus_minus1}),
neuper@37950
   244
	 (*"[| m is_const |] ==> (x + m * v) - v = x + (m - 1) * v"*)
neuper@37950
   245
wneuper@59416
   246
	 Rule.Thm ("subtrahiere_x_minus_plus",TermC.num_str @{thm subtrahiere_x_minus_plus}), 
neuper@37950
   247
	 (*"[| l is_const; m..|] ==> (k - m * n) + l * n = k + (-m + l) * n"*)
wneuper@59416
   248
	 Rule.Thm ("subtrahiere_x_minus1_plus",TermC.num_str @{thm subtrahiere_x_minus1_plus}),
neuper@37950
   249
	 (*"[| l is_const |] ==> (x - v) + l * v = x + (-1 + l) * v"*)
wneuper@59416
   250
	 Rule.Thm ("subtrahiere_x_minus_plus1",TermC.num_str @{thm subtrahiere_x_minus_plus1}),
neuper@37950
   251
	 (*"[| m is_const |] ==> (x - m * v) + v = x + (-m + 1) * v"*)
neuper@37950
   252
wneuper@59416
   253
	 Rule.Thm ("subtrahiere_x_minus_minus",TermC.num_str @{thm subtrahiere_x_minus_minus}), 
neuper@37950
   254
	 (*"[| l is_const; m..|] ==> (k - m * n) - l * n = k + (-m - l) * n"*)
wneuper@59416
   255
	 Rule.Thm ("subtrahiere_x_minus1_minus",TermC.num_str @{thm subtrahiere_x_minus1_minus}),
neuper@37950
   256
	 (*"[| l is_const |] ==> (x - v) - l * v = x + (-1 - l) * v"*)
wneuper@59416
   257
	 Rule.Thm ("subtrahiere_x_minus_minus1",TermC.num_str @{thm subtrahiere_x_minus_minus1}),
neuper@37950
   258
	 (*"[| m is_const |] ==> (x - m * v) - v = x + (-m - 1) * v"*)
neuper@37950
   259
	 
walther@59773
   260
	 Rule.Num_Calc ("Groups.plus_class.plus", (**)eval_binop "#add_"),
walther@59773
   261
	 Rule.Num_Calc ("Groups.minus_class.minus", (**)eval_binop "#subtr_"),
neuper@37950
   262
	 
wneuper@59416
   263
	 (*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen
neuper@37950
   264
           (a+a)+a --> a + 2*a --> 3*a and not (a+a)+a --> 2*a + a *)
wneuper@59416
   265
	 Rule.Thm ("real_mult_2_assoc_r",TermC.num_str @{thm real_mult_2_assoc_r}),
neuper@37950
   266
	 (*"(k + z1) + z1 = k + 2 * z1"*)
wneuper@59416
   267
	 Rule.Thm ("sym_real_mult_2",TermC.num_str (@{thm real_mult_2} RS @{thm sym})),
neuper@37950
   268
	 (*"z1 + z1 = 2 * z1"*)
neuper@37950
   269
wneuper@59416
   270
	 Rule.Thm ("addiere_vor_minus",TermC.num_str @{thm addiere_vor_minus}),
neuper@37950
   271
	 (*"[| l is_const; m is_const |] ==> -(l * v) +  m * v = (-l + m) *v"*)
wneuper@59416
   272
	 Rule.Thm ("addiere_eins_vor_minus",TermC.num_str @{thm addiere_eins_vor_minus}),
neuper@37950
   273
	 (*"[| m is_const |] ==> -  v +  m * v = (-1 + m) * v"*)
wneuper@59416
   274
	 Rule.Thm ("subtrahiere_vor_minus",TermC.num_str @{thm subtrahiere_vor_minus}),
neuper@37950
   275
	 (*"[| l is_const; m is_const |] ==> -(l * v) -  m * v = (-l - m) *v"*)
wneuper@59416
   276
	 Rule.Thm ("subtrahiere_eins_vor_minus",TermC.num_str @{thm subtrahiere_eins_vor_minus})
neuper@37950
   277
	 (*"[| m is_const |] ==> -  v -  m * v = (-1 - m) * v"*)
neuper@37950
   278
	 
wneuper@59416
   279
	 ], scr = Rule.EmptyScr};
neuper@37950
   280
    
neuper@37950
   281
val verschoenere = 
walther@59851
   282
  Rule_Def.Repeat{id = "verschoenere", preconds = [], 
walther@59857
   283
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59852
   284
      erls = Rule_Set.append_rules "erls_verschoenere" Rule_Set.empty 
walther@59773
   285
			[Rule.Num_Calc ("PolyMinus.kleiner", eval_kleiner "")], 
wneuper@59416
   286
      rules = [Rule.Thm ("vorzeichen_minus_weg1",TermC.num_str @{thm vorzeichen_minus_weg1}),
neuper@37950
   287
	       (*"l kleiner 0 ==> a + l * b = a - -l * b"*)
wneuper@59416
   288
	       Rule.Thm ("vorzeichen_minus_weg2",TermC.num_str @{thm vorzeichen_minus_weg2}),
neuper@37950
   289
	       (*"l kleiner 0 ==> a - l * b = a + -l * b"*)
wneuper@59416
   290
	       Rule.Thm ("vorzeichen_minus_weg3",TermC.num_str @{thm vorzeichen_minus_weg3}),
neuper@37950
   291
	       (*"l kleiner 0 ==> k + a - l * b = k + a + -l * b"*)
wneuper@59416
   292
	       Rule.Thm ("vorzeichen_minus_weg4",TermC.num_str @{thm vorzeichen_minus_weg4}),
neuper@37950
   293
	       (*"l kleiner 0 ==> k - a - l * b = k - a + -l * b"*)
neuper@37950
   294
walther@59773
   295
	       Rule.Num_Calc ("Groups.times_class.times", (**)eval_binop "#mult_"),
neuper@37950
   296
wneuper@59416
   297
	       Rule.Thm ("mult_zero_left",TermC.num_str @{thm mult_zero_left}),    
neuper@37950
   298
	       (*"0 * z = 0"*)
wneuper@59416
   299
	       Rule.Thm ("mult_1_left",TermC.num_str @{thm mult_1_left}),     
neuper@37950
   300
	       (*"1 * z = z"*)
wneuper@59416
   301
	       Rule.Thm ("add_0_left",TermC.num_str @{thm add_0_left}),
neuper@37950
   302
	       (*"0 + z = z"*)
wneuper@59416
   303
	       Rule.Thm ("null_minus",TermC.num_str @{thm null_minus}),
neuper@37950
   304
	       (*"0 - a = -a"*)
wneuper@59416
   305
	       Rule.Thm ("vor_minus_mal",TermC.num_str @{thm vor_minus_mal})
neuper@37950
   306
	       (*"- a * b = (-a) * b"*)
neuper@37950
   307
wneuper@59416
   308
	       (*Rule.Thm ("",TermC.num_str @{}),*)
neuper@37950
   309
	       (**)
wneuper@59416
   310
	       ], scr = Rule.EmptyScr} (*end verschoenere*);
neuper@37950
   311
neuper@37950
   312
val klammern_aufloesen = 
walther@59851
   313
  Rule_Def.Repeat{id = "klammern_aufloesen", preconds = [], 
walther@59857
   314
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), srls = Rule_Set.Empty, calc = [], errpatts = [], erls = Rule_Set.Empty, 
wneuper@59416
   315
      rules = [Rule.Thm ("sym_add_assoc",
wneuper@59389
   316
                     TermC.num_str (@{thm add.assoc} RS @{thm sym})),
neuper@37950
   317
	       (*"a + (b + c) = (a + b) + c"*)
wneuper@59416
   318
	       Rule.Thm ("klammer_plus_minus",TermC.num_str @{thm klammer_plus_minus}),
neuper@37950
   319
	       (*"a + (b - c) = (a + b) - c"*)
wneuper@59416
   320
	       Rule.Thm ("klammer_minus_plus",TermC.num_str @{thm klammer_minus_plus}),
neuper@37950
   321
	       (*"a - (b + c) = (a - b) - c"*)
wneuper@59416
   322
	       Rule.Thm ("klammer_minus_minus",TermC.num_str @{thm klammer_minus_minus})
neuper@37950
   323
	       (*"a - (b - c) = (a - b) + c"*)
wneuper@59416
   324
	       ], scr = Rule.EmptyScr};
neuper@37950
   325
neuper@37950
   326
val klammern_ausmultiplizieren = 
walther@59851
   327
  Rule_Def.Repeat{id = "klammern_ausmultiplizieren", preconds = [], 
walther@59857
   328
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), srls = Rule_Set.Empty, calc = [], errpatts = [], erls = Rule_Set.Empty, 
wneuper@59416
   329
      rules = [Rule.Thm ("distrib_right" ,TermC.num_str @{thm distrib_right}),
neuper@37950
   330
	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
wneuper@59416
   331
	       Rule.Thm ("distrib_left",TermC.num_str @{thm distrib_left}),
neuper@37950
   332
	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
neuper@37950
   333
	       
wneuper@59416
   334
	       Rule.Thm ("klammer_mult_minus",TermC.num_str @{thm klammer_mult_minus}),
neuper@37950
   335
	       (*"a * (b - c) = a * b - a * c"*)
wneuper@59416
   336
	       Rule.Thm ("klammer_minus_mult",TermC.num_str @{thm klammer_minus_mult})
neuper@37950
   337
	       (*"(b - c) * a = b * a - c * a"*)
neuper@37950
   338
wneuper@59416
   339
	       (*Rule.Thm ("",TermC.num_str @{}),
neuper@37950
   340
	       (*""*)*)
wneuper@59416
   341
	       ], scr = Rule.EmptyScr};
neuper@37950
   342
neuper@37950
   343
val ordne_monome = 
walther@59851
   344
  Rule_Def.Repeat{id = "ordne_monome", preconds = [], 
walther@59857
   345
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), srls = Rule_Set.Empty, calc = [], errpatts = [], 
walther@59852
   346
      erls = Rule_Set.append_rules "erls_ordne_monome" Rule_Set.empty
walther@59773
   347
	       [Rule.Num_Calc ("PolyMinus.kleiner", eval_kleiner ""),
walther@59773
   348
		Rule.Num_Calc ("Prog_Expr.is'_atom", Prog_Expr.eval_is_atom "")
neuper@37950
   349
		], 
wneuper@59416
   350
      rules = [Rule.Thm ("tausche_mal",TermC.num_str @{thm tausche_mal}),
neuper@37950
   351
	       (*"[| b is_atom; a kleiner b  |] ==> (b * a) = (a * b)"*)
wneuper@59416
   352
	       Rule.Thm ("tausche_vor_mal",TermC.num_str @{thm tausche_vor_mal}),
neuper@37950
   353
	       (*"[| b is_atom; a kleiner b  |] ==> (-b * a) = (-a * b)"*)
wneuper@59416
   354
	       Rule.Thm ("tausche_mal_mal",TermC.num_str @{thm tausche_mal_mal}),
neuper@37950
   355
	       (*"[| c is_atom; b kleiner c  |] ==> (a * c * b) = (a * b *c)"*)
wneuper@59416
   356
	       Rule.Thm ("x_quadrat",TermC.num_str @{thm x_quadrat})
neuper@37950
   357
	       (*"(x * a) * a = x * a ^^^ 2"*)
neuper@37950
   358
wneuper@59416
   359
	       (*Rule.Thm ("",TermC.num_str @{}),
neuper@37950
   360
	       (*""*)*)
wneuper@59416
   361
	       ], scr = Rule.EmptyScr};
neuper@37950
   362
neuper@37950
   363
neuper@37950
   364
val rls_p_33 = 
walther@59852
   365
    Rule_Set.append_rules "rls_p_33" Rule_Set.empty
wneuper@59416
   366
	       [Rule.Rls_ ordne_alphabetisch,
wneuper@59416
   367
		Rule.Rls_ fasse_zusammen,
wneuper@59416
   368
		Rule.Rls_ verschoenere
neuper@37950
   369
		];
neuper@37950
   370
val rls_p_34 = 
walther@59852
   371
    Rule_Set.append_rules "rls_p_34" Rule_Set.empty
wneuper@59416
   372
	       [Rule.Rls_ klammern_aufloesen,
wneuper@59416
   373
		Rule.Rls_ ordne_alphabetisch,
wneuper@59416
   374
		Rule.Rls_ fasse_zusammen,
wneuper@59416
   375
		Rule.Rls_ verschoenere
neuper@37950
   376
		];
neuper@37950
   377
val rechnen = 
walther@59852
   378
    Rule_Set.append_rules "rechnen" Rule_Set.empty
walther@59773
   379
	       [Rule.Num_Calc ("Groups.times_class.times", (**)eval_binop "#mult_"),
walther@59773
   380
		Rule.Num_Calc ("Groups.plus_class.plus", (**)eval_binop "#add_"),
walther@59773
   381
		Rule.Num_Calc ("Groups.minus_class.minus", (**)eval_binop "#subtr_")
neuper@37950
   382
		];
wneuper@59472
   383
\<close>
wneuper@59472
   384
setup \<open>KEStore_Elems.add_rlss 
s1210629013@55444
   385
  [("ordne_alphabetisch", (Context.theory_name @{theory}, prep_rls' ordne_alphabetisch)), 
s1210629013@55444
   386
  ("fasse_zusammen", (Context.theory_name @{theory}, prep_rls' fasse_zusammen)), 
s1210629013@55444
   387
  ("verschoenere", (Context.theory_name @{theory}, prep_rls' verschoenere)), 
s1210629013@55444
   388
  ("ordne_monome", (Context.theory_name @{theory}, prep_rls' ordne_monome)), 
s1210629013@55444
   389
  ("klammern_aufloesen", (Context.theory_name @{theory}, prep_rls' klammern_aufloesen)), 
neuper@52125
   390
  ("klammern_ausmultiplizieren",
wneuper@59472
   391
    (Context.theory_name @{theory}, prep_rls' klammern_ausmultiplizieren))]\<close>
neuper@37950
   392
neuper@37950
   393
(** problems **)
wneuper@59472
   394
setup \<open>KEStore_Elems.add_pbts
wneuper@59406
   395
  [(Specify.prep_pbt thy "pbl_vereinf_poly" [] Celem.e_pblID
walther@59851
   396
      (["polynom","vereinfachen"], [], Rule_Set.Empty, NONE, [])),
wneuper@59406
   397
    (Specify.prep_pbt thy "pbl_vereinf_poly_minus" [] Celem.e_pblID
s1210629013@55339
   398
      (["plus_minus","polynom","vereinfachen"],
s1210629013@55339
   399
        [("#Given", ["Term t_t"]),
s1210629013@55339
   400
          ("#Where", ["t_t is_polyexp",
s1210629013@55339
   401
            "Not (matchsub (?a + (?b + ?c)) t_t | " ^
s1210629013@55339
   402
            "     matchsub (?a + (?b - ?c)) t_t | " ^
s1210629013@55339
   403
            "     matchsub (?a - (?b + ?c)) t_t | " ^
s1210629013@55339
   404
            "     matchsub (?a + (?b - ?c)) t_t )",
s1210629013@55339
   405
            "Not (matchsub (?a * (?b + ?c)) t_t | " ^
s1210629013@55339
   406
            "     matchsub (?a * (?b - ?c)) t_t | " ^
s1210629013@55339
   407
            "     matchsub ((?b + ?c) * ?a) t_t | " ^
s1210629013@55339
   408
            "     matchsub ((?b - ?c) * ?a) t_t )"]),
s1210629013@55339
   409
          ("#Find", ["normalform n_n"])],
walther@59852
   410
        Rule_Set.append_rules "prls_pbl_vereinf_poly" Rule_Set.empty 
walther@59773
   411
	        [Rule.Num_Calc ("Poly.is'_polyexp", eval_is_polyexp ""),
walther@59773
   412
	          Rule.Num_Calc ("Prog_Expr.matchsub", Prog_Expr.eval_matchsub ""),
wneuper@59416
   413
	          Rule.Thm ("or_true", TermC.num_str @{thm or_true}),
s1210629013@55339
   414
            (*"(?a | True) = True"*)
wneuper@59416
   415
            Rule.Thm ("or_false", TermC.num_str @{thm or_false}),
s1210629013@55339
   416
            (*"(?a | False) = ?a"*)
wneuper@59416
   417
            Rule.Thm ("not_true",TermC.num_str @{thm not_true}),
s1210629013@55339
   418
            (*"(~ True) = False"*)
wneuper@59416
   419
            Rule.Thm ("not_false",TermC.num_str @{thm not_false})
s1210629013@55339
   420
            (*"(~ False) = True"*)], 
s1210629013@55339
   421
       SOME "Vereinfache t_t", [["simplification","for_polynomials","with_minus"]])),
wneuper@59406
   422
    (Specify.prep_pbt thy "pbl_vereinf_poly_klammer" [] Celem.e_pblID
s1210629013@55339
   423
      (["klammer","polynom","vereinfachen"],
s1210629013@55339
   424
        [("#Given" ,["Term t_t"]),
s1210629013@55339
   425
          ("#Where" ,["t_t is_polyexp",
s1210629013@55339
   426
            "Not (matchsub (?a * (?b + ?c)) t_t | " ^
s1210629013@55339
   427
            "     matchsub (?a * (?b - ?c)) t_t | " ^
s1210629013@55339
   428
            "     matchsub ((?b + ?c) * ?a) t_t | " ^
s1210629013@55339
   429
            "     matchsub ((?b - ?c) * ?a) t_t )"]),
s1210629013@55339
   430
          ("#Find"  ,["normalform n_n"])],
walther@59852
   431
        Rule_Set.append_rules "prls_pbl_vereinf_poly_klammer" Rule_Set.empty
walther@59773
   432
          [Rule.Num_Calc ("Poly.is'_polyexp", eval_is_polyexp ""),
walther@59773
   433
	           Rule.Num_Calc ("Prog_Expr.matchsub", Prog_Expr.eval_matchsub ""),
wneuper@59416
   434
             Rule.Thm ("or_true", TermC.num_str @{thm or_true}),
s1210629013@55339
   435
             (*"(?a | True) = True"*)
wneuper@59416
   436
             Rule.Thm ("or_false", TermC.num_str @{thm or_false}),
s1210629013@55339
   437
             (*"(?a | False) = ?a"*)
wneuper@59416
   438
             Rule.Thm ("not_true",TermC.num_str @{thm not_true}),
s1210629013@55339
   439
             (*"(~ True) = False"*)
wneuper@59416
   440
             Rule.Thm ("not_false",TermC.num_str @{thm not_false})
s1210629013@55339
   441
             (*"(~ False) = True"*)], 
s1210629013@55339
   442
        SOME "Vereinfache t_t", 
s1210629013@55339
   443
        [["simplification","for_polynomials","with_parentheses"]])),
wneuper@59406
   444
    (Specify.prep_pbt thy "pbl_vereinf_poly_klammer_mal" [] Celem.e_pblID
s1210629013@55339
   445
      (["binom_klammer","polynom","vereinfachen"],
s1210629013@55339
   446
        [("#Given", ["Term t_t"]),
s1210629013@55339
   447
          ("#Where", ["t_t is_polyexp"]),
s1210629013@55339
   448
          ("#Find", ["normalform n_n"])],
walther@59852
   449
        Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)
walther@59773
   450
			      Rule.Num_Calc ("Poly.is'_polyexp", eval_is_polyexp "")], 
s1210629013@55339
   451
        SOME "Vereinfache t_t", 
s1210629013@55339
   452
        [["simplification","for_polynomials","with_parentheses_mult"]])),
walther@59851
   453
    (Specify.prep_pbt thy "pbl_probe" [] Celem.e_pblID (["probe"], [], Rule_Set.Empty, NONE, [])),
wneuper@59406
   454
    (Specify.prep_pbt thy "pbl_probe_poly" [] Celem.e_pblID
s1210629013@55339
   455
      (["polynom","probe"],
s1210629013@55339
   456
        [("#Given", ["Pruefe e_e", "mitWert w_w"]),
s1210629013@55339
   457
          ("#Where", ["e_e is_polyexp"]),
s1210629013@55339
   458
          ("#Find", ["Geprueft p_p"])],
walther@59852
   459
        Rule_Set.append_rules "prls_pbl_probe_poly" Rule_Set.empty [(*for preds in where_*)
walther@59773
   460
		      Rule.Num_Calc ("Poly.is'_polyexp", eval_is_polyexp "")], 
s1210629013@55339
   461
        SOME "Probe e_e w_w", 
s1210629013@55339
   462
        [["probe","fuer_polynom"]])),
wneuper@59406
   463
    (Specify.prep_pbt thy "pbl_probe_bruch" [] Celem.e_pblID
s1210629013@55339
   464
      (["bruch","probe"],
s1210629013@55339
   465
        [("#Given" ,["Pruefe e_e", "mitWert w_w"]),
s1210629013@55339
   466
          ("#Where" ,["e_e is_ratpolyexp"]),
s1210629013@55339
   467
          ("#Find"  ,["Geprueft p_p"])],
walther@59852
   468
        Rule_Set.append_rules "prls_pbl_probe_bruch" Rule_Set.empty [(*for preds in where_*)
walther@59773
   469
		      Rule.Num_Calc ("Rational.is'_ratpolyexp", eval_is_ratpolyexp "")], 
wneuper@59472
   470
        SOME "Probe e_e w_w", [["probe","fuer_bruch"]]))]\<close>
neuper@37950
   471
neuper@37950
   472
(** methods **)
wneuper@59545
   473
wneuper@59504
   474
partial_function (tailrec) simplify :: "real \<Rightarrow> real"
wneuper@59504
   475
  where
walther@59635
   476
"simplify t_t = (
walther@59635
   477
  (Repeat(
walther@59637
   478
    (Try (Rewrite_Set ''ordne_alphabetisch'')) #>
walther@59637
   479
    (Try (Rewrite_Set ''fasse_zusammen'')) #>
walther@59635
   480
    (Try (Rewrite_Set ''verschoenere'')))
walther@59635
   481
  ) t_t)"
wneuper@59472
   482
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   483
    [Specify.prep_met thy "met_simp_poly_minus" [] Celem.e_metID
s1210629013@55373
   484
	    (["simplification","for_polynomials","with_minus"],
s1210629013@55373
   485
	      [("#Given" ,["Term t_t"]),
s1210629013@55373
   486
	        ("#Where" ,["t_t is_polyexp",
s1210629013@55373
   487
	            "Not (matchsub (?a + (?b + ?c)) t_t | " ^
s1210629013@55373
   488
	            "     matchsub (?a + (?b - ?c)) t_t | " ^
s1210629013@55373
   489
	            "     matchsub (?a - (?b + ?c)) t_t | " ^
s1210629013@55373
   490
	            "     matchsub (?a + (?b - ?c)) t_t )"]),
s1210629013@55373
   491
	        ("#Find"  ,["normalform n_n"])],
walther@59852
   492
	      {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty,
walther@59852
   493
	        prls = Rule_Set.append_rules "prls_met_simp_poly_minus" Rule_Set.empty 
walther@59773
   494
				      [Rule.Num_Calc ("Poly.is'_polyexp", eval_is_polyexp ""),
walther@59773
   495
				        Rule.Num_Calc ("Prog_Expr.matchsub", Prog_Expr.eval_matchsub ""),
wneuper@59416
   496
				        Rule.Thm ("and_true",TermC.num_str @{thm and_true}),
s1210629013@55373
   497
                (*"(?a & True) = ?a"*)
wneuper@59416
   498
                Rule.Thm ("and_false",TermC.num_str @{thm and_false}),
s1210629013@55373
   499
                (*"(?a & False) = False"*)
wneuper@59416
   500
                Rule.Thm ("not_true",TermC.num_str @{thm not_true}),
s1210629013@55373
   501
                (*"(~ True) = False"*)
wneuper@59416
   502
                Rule.Thm ("not_false",TermC.num_str @{thm not_false})
s1210629013@55373
   503
                (*"(~ False) = True"*)],
walther@59852
   504
          crls = Rule_Set.empty, errpats = [], nrls = rls_p_33},
wneuper@59551
   505
          @{thm simplify.simps})]
wneuper@59473
   506
\<close>
wneuper@59545
   507
wneuper@59504
   508
partial_function (tailrec) simplify2 :: "real \<Rightarrow> real"
wneuper@59504
   509
  where
walther@59635
   510
"simplify2 t_t = (
walther@59635
   511
  (Repeat(
walther@59637
   512
    (Try (Rewrite_Set ''klammern_aufloesen'')) #>
walther@59637
   513
    (Try (Rewrite_Set ''ordne_alphabetisch'')) #>
walther@59637
   514
    (Try (Rewrite_Set ''fasse_zusammen'')) #>
walther@59635
   515
    (Try (Rewrite_Set ''verschoenere'')))
walther@59635
   516
  ) t_t)"
wneuper@59473
   517
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   518
    [Specify.prep_met thy "met_simp_poly_parenth" [] Celem.e_metID
s1210629013@55373
   519
	    (["simplification","for_polynomials","with_parentheses"],
s1210629013@55373
   520
	      [("#Given" ,["Term t_t"]),
s1210629013@55373
   521
	        ("#Where" ,["t_t is_polyexp"]),
s1210629013@55373
   522
	        ("#Find"  ,["normalform n_n"])],
walther@59852
   523
	      {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, 
walther@59852
   524
	        prls = Rule_Set.append_rules "simplification_for_polynomials_prls" Rule_Set.empty 
walther@59773
   525
				    [(*for preds in where_*) Rule.Num_Calc("Poly.is'_polyexp", eval_is_polyexp"")],
walther@59852
   526
				  crls = Rule_Set.empty, errpats = [], nrls = rls_p_34},
wneuper@59551
   527
				@{thm simplify2.simps})]
wneuper@59473
   528
\<close>
wneuper@59545
   529
wneuper@59504
   530
partial_function (tailrec) simplify3 :: "real \<Rightarrow> real"
wneuper@59504
   531
  where
walther@59635
   532
"simplify3 t_t = (
walther@59635
   533
  (Repeat(
walther@59637
   534
    (Try (Rewrite_Set ''klammern_ausmultiplizieren'')) #>
walther@59637
   535
    (Try (Rewrite_Set ''discard_parentheses'')) #>
walther@59637
   536
    (Try (Rewrite_Set ''ordne_monome'')) #>
walther@59637
   537
    (Try (Rewrite_Set ''klammern_aufloesen'')) #>
walther@59637
   538
    (Try (Rewrite_Set ''ordne_alphabetisch'')) #>
walther@59637
   539
    (Try (Rewrite_Set ''fasse_zusammen'')) #>
walther@59635
   540
    (Try (Rewrite_Set ''verschoenere'')))
walther@59635
   541
  ) t_t)"
wneuper@59473
   542
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   543
    [Specify.prep_met thy "met_simp_poly_parenth_mult" [] Celem.e_metID
s1210629013@55373
   544
	    (["simplification","for_polynomials","with_parentheses_mult"],
s1210629013@55373
   545
	      [("#Given" ,["Term t_t"]), ("#Where" ,["t_t is_polyexp"]), ("#Find"  ,["normalform n_n"])],
walther@59852
   546
	        {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, 
walther@59852
   547
	          prls = Rule_Set.append_rules "simplification_for_polynomials_prls" Rule_Set.empty 
walther@59773
   548
				      [(*for preds in where_*) Rule.Num_Calc("Poly.is'_polyexp", eval_is_polyexp"")],
walther@59852
   549
				    crls = Rule_Set.empty, errpats = [], nrls = rls_p_34},
wneuper@59551
   550
				  @{thm simplify3.simps})]
wneuper@59473
   551
\<close>
wneuper@59473
   552
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   553
    [Specify.prep_met thy "met_probe" [] Celem.e_metID
s1210629013@55373
   554
	    (["probe"], [],
walther@59852
   555
	      {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.Empty, crls = Rule_Set.empty,
walther@59851
   556
	        errpats = [], nrls = Rule_Set.Empty}, 
wneuper@59545
   557
	      @{thm refl})]
wneuper@59473
   558
\<close>
wneuper@59545
   559
wneuper@59504
   560
partial_function (tailrec) mache_probe :: "bool \<Rightarrow> bool list \<Rightarrow> bool"
wneuper@59504
   561
  where
walther@59635
   562
"mache_probe e_e w_w = (
walther@59635
   563
  let
walther@59635
   564
     e_e = Take e_e;
walther@59635
   565
     e_e = Substitute w_w e_e
walther@59635
   566
  in (
walther@59635
   567
    Repeat (
walther@59637
   568
      (Try (Repeat (Calculate ''TIMES''))) #>
walther@59637
   569
      (Try (Repeat (Calculate ''PLUS'' ))) #>
walther@59635
   570
      (Try (Repeat (Calculate ''MINUS''))))
walther@59635
   571
    ) e_e)"
wneuper@59473
   572
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   573
    [Specify.prep_met thy "met_probe_poly" [] Celem.e_metID
s1210629013@55373
   574
	    (["probe","fuer_polynom"],
s1210629013@55373
   575
	      [("#Given" ,["Pruefe e_e", "mitWert w_w"]),
s1210629013@55373
   576
	        ("#Where" ,["e_e is_polyexp"]),
s1210629013@55373
   577
	        ("#Find"  ,["Geprueft p_p"])],
walther@59852
   578
	      {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, 
walther@59852
   579
	        prls = Rule_Set.append_rules "prls_met_probe_bruch" Rule_Set.empty
walther@59773
   580
	            [(*for preds in where_*) Rule.Num_Calc ("Rational.is'_ratpolyexp", eval_is_ratpolyexp "")], 
walther@59852
   581
	        crls = Rule_Set.empty, errpats = [], nrls = rechnen}, 
wneuper@59551
   582
	      @{thm mache_probe.simps})]
wneuper@59473
   583
\<close>
wneuper@59473
   584
setup \<open>KEStore_Elems.add_mets
wneuper@59473
   585
    [Specify.prep_met thy "met_probe_bruch" [] Celem.e_metID
s1210629013@55373
   586
	    (["probe","fuer_bruch"],
s1210629013@55373
   587
	      [("#Given" ,["Pruefe e_e", "mitWert w_w"]),
s1210629013@55373
   588
	        ("#Where" ,["e_e is_ratpolyexp"]),
s1210629013@55373
   589
	        ("#Find"  ,["Geprueft p_p"])],
walther@59852
   590
	      {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, 
walther@59852
   591
	        prls = Rule_Set.append_rules "prls_met_probe_bruch" Rule_Set.empty
walther@59773
   592
	            [(*for preds in where_*) Rule.Num_Calc ("Rational.is'_ratpolyexp", eval_is_ratpolyexp "")], 
walther@59852
   593
	        crls = Rule_Set.empty, errpats = [], nrls = Rule_Set.Empty}, 
wneuper@59545
   594
	      @{thm refl})]
wneuper@59472
   595
\<close>
neuper@37906
   596
neuper@37906
   597
end
neuper@37906
   598