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