neuper@37906: (* attempts to perserve binary minus as wanted by Austrian teachers neuper@37906: WN071207 neuper@37906: (c) due to copyright terms neuper@37906: *) neuper@37906: neuper@37950: theory PolyMinus imports (*Poly// due to "is_ratpolyexp" in...*) Rational begin neuper@37906: neuper@37906: consts neuper@37906: neuper@37906: (*predicates for conditions in rewriting*) neuper@37906: kleiner :: "['a, 'a] => bool" ("_ kleiner _") walther@60278: ist_monom :: "'a => bool" ("_ ist'_monom") neuper@37906: neuper@37906: (*the CAS-command*) neuper@37906: Probe :: "[bool, bool list] => bool" neuper@37906: (*"Probe (3*a+2*b+a = 4*a+2*b) [a=1,b=2]"*) neuper@37906: neuper@37906: (*descriptions for the pbl and met*) neuper@37950: Pruefe :: "bool => una" neuper@37950: mitWert :: "bool list => tobooll" neuper@37950: Geprueft :: "bool => una" neuper@37906: neuper@52148: axiomatization where neuper@37906: neuper@52148: null_minus: "0 - a = -a" and neuper@52148: vor_minus_mal: "- a * b = (-a) * b" and neuper@37906: neuper@37906: (*commute with invariant (a.b).c -association*) neuper@37980: tausche_plus: "[| b ist_monom; a kleiner b |] ==> neuper@52148: (b + a) = (a + b)" and neuper@37980: tausche_minus: "[| b ist_monom; a kleiner b |] ==> neuper@52148: (b - a) = (-a + b)" and neuper@37980: tausche_vor_plus: "[| b ist_monom; a kleiner b |] ==> neuper@52148: (- b + a) = (a - b)" and neuper@37980: tausche_vor_minus: "[| b ist_monom; a kleiner b |] ==> neuper@52148: (- b - a) = (-a - b)" and walther@60269: (*Ambiguous input\<^here> produces 3 parse trees -----------------------------\\*) neuper@52148: tausche_plus_plus: "b kleiner c ==> (a + c + b) = (a + b + c)" and neuper@52148: tausche_plus_minus: "b kleiner c ==> (a + c - b) = (a - b + c)" and neuper@52148: tausche_minus_plus: "b kleiner c ==> (a - c + b) = (a + b - c)" and neuper@52148: tausche_minus_minus: "b kleiner c ==> (a - c - b) = (a - b - c)" and walther@60269: (*Ambiguous input\<^here> produces 3 parse trees -----------------------------//*) neuper@37906: neuper@37906: (*commute with invariant (a.b).c -association*) neuper@37980: tausche_mal: "[| b is_atom; a kleiner b |] ==> neuper@52148: (b * a) = (a * b)" and neuper@37980: tausche_vor_mal: "[| b is_atom; a kleiner b |] ==> neuper@52148: (-b * a) = (-a * b)" and neuper@37980: tausche_mal_mal: "[| c is_atom; b kleiner c |] ==> neuper@52148: (x * c * b) = (x * b * c)" and walther@60242: x_quadrat: "(x * a) * a = x * a \ 2" and neuper@37906: neuper@37906: walther@60385: subtrahiere: "[| l is_num; m is_num |] ==> neuper@52148: m * v - l * v = (m - l) * v" and walther@60385: subtrahiere_von_1: "[| l is_num |] ==> neuper@52148: v - l * v = (1 - l) * v" and walther@60385: subtrahiere_1: "[| l is_num; m is_num |] ==> neuper@52148: m * v - v = (m - 1) * v" and neuper@37906: walther@60385: subtrahiere_x_plus_minus: "[| l is_num; m is_num |] ==> neuper@52148: (x + m * v) - l * v = x + (m - l) * v" and walther@60385: subtrahiere_x_plus1_minus: "[| l is_num |] ==> neuper@52148: (x + v) - l * v = x + (1 - l) * v" and walther@60385: subtrahiere_x_plus_minus1: "[| m is_num |] ==> neuper@52148: (x + m * v) - v = x + (m - 1) * v" and neuper@37906: walther@60385: subtrahiere_x_minus_plus: "[| l is_num; m is_num |] ==> neuper@52148: (x - m * v) + l * v = x + (-m + l) * v" and walther@60385: subtrahiere_x_minus1_plus: "[| l is_num |] ==> neuper@52148: (x - v) + l * v = x + (-1 + l) * v" and walther@60385: subtrahiere_x_minus_plus1: "[| m is_num |] ==> neuper@52148: (x - m * v) + v = x + (-m + 1) * v" and neuper@37906: walther@60385: subtrahiere_x_minus_minus: "[| l is_num; m is_num |] ==> neuper@52148: (x - m * v) - l * v = x + (-m - l) * v" and walther@60385: subtrahiere_x_minus1_minus:"[| l is_num |] ==> neuper@52148: (x - v) - l * v = x + (-1 - l) * v" and walther@60385: subtrahiere_x_minus_minus1:"[| m is_num |] ==> neuper@52148: (x - m * v) - v = x + (-m - 1) * v" and neuper@37906: neuper@37906: walther@60385: addiere_vor_minus: "[| l is_num; m is_num |] ==> neuper@52148: - (l * v) + m * v = (-l + m) * v" and walther@60385: addiere_eins_vor_minus: "[| m is_num |] ==> neuper@52148: - v + m * v = (-1 + m) * v" and walther@60385: subtrahiere_vor_minus: "[| l is_num; m is_num |] ==> neuper@52148: - (l * v) - m * v = (-l - m) * v" and walther@60385: subtrahiere_eins_vor_minus:"[| m is_num |] ==> neuper@52148: - v - m * v = (-1 - m) * v" and neuper@37906: walther@60269: (*Ambiguous input\<^here> produces 3 parse trees -----------------------------\\*) neuper@52148: vorzeichen_minus_weg1: "l kleiner 0 ==> a + l * b = a - -1*l * b" and neuper@52148: vorzeichen_minus_weg2: "l kleiner 0 ==> a - l * b = a + -1*l * b" and neuper@52148: vorzeichen_minus_weg3: "l kleiner 0 ==> k + a - l * b = k + a + -1*l * b" and neuper@52148: vorzeichen_minus_weg4: "l kleiner 0 ==> k - a - l * b = k - a + -1*l * b" and walther@60269: (*Ambiguous input\<^here> produces 3 parse trees -----------------------------//*) neuper@37906: walther@59877: (*klammer_plus_plus = (add.assoc RS sym)*) neuper@52148: klammer_plus_minus: "a + (b - c) = (a + b) - c" and neuper@52148: klammer_minus_plus: "a - (b + c) = (a - b) - c" and neuper@52148: klammer_minus_minus: "a - (b - c) = (a - b) + c" and neuper@37906: neuper@52148: klammer_mult_minus: "a * (b - c) = a * b - a * c" and neuper@37980: klammer_minus_mult: "(b - c) * a = b * a - c * a" neuper@37906: wneuper@59472: ML \ neuper@37950: (** eval functions **) neuper@37906: neuper@37950: (*. get the identifier from specific monomials; see fun ist_monom .*) walther@60351: fun Free_to_string (Free (str, _)) = str walther@60351: | Free_to_string t = walther@60351: if TermC.is_num t then TermC.string_of_num t else "|||||||||||||"(*the "largest" string*); walther@60351: walther@60351: (* quick and dirty solution just before a field test *) walther@60351: fun identifier (Free (id,_)) = id (* _a_ *) walther@60351: | identifier (Const (\<^const_name>\times\, _) $ t1 $ t2) = (* 2*_a_, a*_b_, 3*a*_b_ *) walther@60351: if TermC.is_atom t2 walther@60351: then Free_to_string t2 walther@60351: else walther@60351: (case t1 of walther@60351: Const (\<^const_name>\times\, _) $ num $ t1' => (* 3*_a_ \ 2 *) walther@60351: if TermC.is_atom num andalso TermC.is_atom t1' then Free_to_string t2 walther@60351: else "|||||||||||||" walther@60351: | _ => walther@60351: (case t2 of wenzelm@60405: Const (\<^const_name>\realpow\, _) $ base $ exp => walther@60351: if TermC.is_atom base andalso TermC.is_atom exp then walther@60351: if TermC.is_num base then "|||||||||||||" walther@60351: else Free_to_string base walther@60351: else "|||||||||||||" walther@60351: | _ => "|||||||||||||")) wenzelm@60405: | identifier (Const (\<^const_name>\realpow\, _) $ base $ exp) = (* _a_\2, _3_^2 *) walther@60351: if TermC.is_atom base andalso TermC.is_atom exp then Free_to_string base neuper@37950: else "|||||||||||||" walther@60351: | identifier t = (* 12 *) walther@60351: if TermC.is_num t walther@60351: then TermC.string_of_num t walther@60351: else "|||||||||||||" (*the "largest" string*); neuper@37950: neuper@37950: (*("kleiner", ("PolyMinus.kleiner", eval_kleiner ""))*) neuper@37950: (* order "by alphabet" w.r.t. var: num < (var | num*var) > (var*var | ..) *) walther@60335: fun eval_kleiner _ _ (p as (Const (\<^const_name>\PolyMinus.kleiner\,_) $ a $ b)) _ = walther@60325: if TermC.is_num b then walther@60325: if TermC.is_num a then (*123 kleiner 32 = True !!!*) walther@60325: if TermC.num_of_term a < TermC.num_of_term b then walther@60325: SOME ((UnparseC.term p) ^ " = True", walther@60325: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@60325: else SOME ((UnparseC.term p) ^ " = False", walther@60325: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) walther@60325: else (* -1 * -2 kleiner 0 *) walther@60325: SOME ((UnparseC.term p) ^ " = False", walther@60325: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) neuper@37950: else walther@60325: if identifier a < identifier b then walther@60325: SOME ((UnparseC.term p) ^ " = True", walther@60325: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@60325: else SOME ((UnparseC.term p) ^ " = False", walther@60325: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) neuper@37950: | eval_kleiner _ _ _ _ = NONE; neuper@37950: walther@60325: fun ist_monom t = walther@60325: if TermC.is_atom t then true walther@60325: else walther@60325: case t of walther@60335: Const (\<^const_name>\Groups.times_class.times\, _) $ t1 $ t2 => walther@60325: ist_monom t1 andalso ist_monom t2 wenzelm@60405: | Const (\<^const_name>\realpow\, _) $ t1 $ t2 => walther@60325: ist_monom t1 andalso ist_monom t2 walther@60325: | _ => false neuper@37950: neuper@37950: (* is this a univariate monomial ? *) walther@60278: (*("ist_monom", ("PolyMinus.ist_monom", eval_ist_monom ""))*) walther@60335: fun eval_ist_monom _ _ (p as (Const (\<^const_name>\PolyMinus.ist_monom\,_) $ a)) _ = neuper@37950: if ist_monom a then walther@59868: SOME ((UnparseC.term p) ^ " = True", wneuper@59390: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@59868: else SOME ((UnparseC.term p) ^ " = False", wneuper@59390: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) neuper@37950: | eval_ist_monom _ _ _ _ = NONE; neuper@37950: neuper@37950: neuper@37950: (** rewrite order **) neuper@37950: neuper@37950: (** rulesets **) neuper@37950: neuper@37950: val erls_ordne_alphabetisch = walther@60358: Rule_Set.append_rules "erls_ordne_alphabetisch" Rule_Set.empty [ walther@60358: \<^rule_eval>\kleiner\ (eval_kleiner ""), walther@60358: \<^rule_eval>\ist_monom\ (eval_ist_monom "")]; neuper@37950: neuper@37950: val ordne_alphabetisch = walther@59851: Rule_Def.Repeat{id = "ordne_alphabetisch", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: erls = erls_ordne_alphabetisch, walther@60358: rules = [ walther@60358: \<^rule_thm>\tausche_plus\, (*"b kleiner a ==> (b + a) = (a + b)"*) walther@60358: \<^rule_thm>\tausche_minus\, (*"b kleiner a ==> (b - a) = (-a + b)"*) walther@60358: \<^rule_thm>\tausche_vor_plus\, (*"[| b ist_monom; a kleiner b |] ==> (- b + a) = (a - b)"*) walther@60358: \<^rule_thm>\tausche_vor_minus\, (*"[| b ist_monom; a kleiner b |] ==> (- b - a) = (-a - b)"*) walther@60358: \<^rule_thm>\tausche_plus_plus\, (*"c kleiner b ==> (a + c + b) = (a + b + c)"*) walther@60358: \<^rule_thm>\tausche_plus_minus\, (*"c kleiner b ==> (a + c - b) = (a - b + c)"*) walther@60358: \<^rule_thm>\tausche_minus_plus\, (*"c kleiner b ==> (a - c + b) = (a + b - c)"*) walther@60358: \<^rule_thm>\tausche_minus_minus\], (*"c kleiner b ==> (a - c - b) = (a - b - c)"*) walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val fasse_zusammen = walther@60358: Rule_Def.Repeat{id = "fasse_zusammen", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Rule_Set.append_rules "erls_fasse_zusammen" Rule_Set.empty walther@60385: [\<^rule_eval>\Prog_Expr.is_num\ (Prog_Expr.eval_is_num "#is_num_")], walther@60358: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60385: \<^rule_thm>\real_num_collect\, (*"[| l is_num; m is_num |]==>l * n + m * n = (l + m) * n"*) walther@60385: \<^rule_thm>\real_num_collect_assoc_r\, (*"[| l is_num; m..|] ==> (k + m * n) + l * n = k + (l + m)*n"*) walther@60385: \<^rule_thm>\real_one_collect\, (*"m is_num ==> n + m * n = (1 + m) * n"*) walther@60385: \<^rule_thm>\real_one_collect_assoc_r\, (*"m is_num ==> (k + n) + m * n = k + (m + 1) * n"*) neuper@37950: walther@60385: \<^rule_thm>\subtrahiere\, (*"[| l is_num; m is_num |] ==> m * v - l * v = (m - l) * v"*) walther@60385: \<^rule_thm>\subtrahiere_von_1\, (*"[| l is_num |] ==> v - l * v = (1 - l) * v"*) walther@60385: \<^rule_thm>\subtrahiere_1\, (*"[| l is_num; m is_num |] ==> m * v - v = (m - 1) * v"*) neuper@37950: walther@60385: \<^rule_thm>\subtrahiere_x_plus_minus\, (*"[| l is_num; m..|] ==> (k + m * n) - l * n = k + ( m - l) * n"*) walther@60385: \<^rule_thm>\subtrahiere_x_plus1_minus\, (*"[| l is_num |] ==> (x + v) - l * v = x + (1 - l) * v"*) walther@60385: \<^rule_thm>\subtrahiere_x_plus_minus1\, (*"[| m is_num |] ==> (x + m * v) - v = x + (m - 1) * v"*) neuper@37950: walther@60385: \<^rule_thm>\subtrahiere_x_minus_plus\, (*"[| l is_num; m..|] ==> (k - m * n) + l * n = k + (-m + l) * n"*) walther@60385: \<^rule_thm>\subtrahiere_x_minus1_plus\, (*"[| l is_num |] ==> (x - v) + l * v = x + (-1 + l) * v"*) walther@60385: \<^rule_thm>\subtrahiere_x_minus_plus1\, (*"[| m is_num |] ==> (x - m * v) + v = x + (-m + 1) * v"*) neuper@37950: walther@60385: \<^rule_thm>\subtrahiere_x_minus_minus\, (*"[| l is_num; m..|] ==> (k - m * n) - l * n = k + (-m - l) * n"*) walther@60385: \<^rule_thm>\subtrahiere_x_minus1_minus\, (*"[| l is_num |] ==> (x - v) - l * v = x + (-1 - l) * v"*) walther@60385: \<^rule_thm>\subtrahiere_x_minus_minus1\, (*"[| m is_num |] ==> (x - m * v) - v = x + (-m - 1) * v"*) neuper@37950: walther@60358: \<^rule_eval>\plus\ (**)(eval_binop "#add_"), walther@60358: \<^rule_eval>\minus\ (**)(eval_binop "#subtr_"), neuper@37950: walther@60358: (*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen walther@60358: (a+a)+a --> a + 2*a --> 3*a and not (a+a)+a --> 2*a + a *) walther@60358: \<^rule_thm>\real_mult_2_assoc_r\, (*"(k + z1) + z1 = k + 2 * z1"*) walther@60358: \<^rule_thm_sym>\real_mult_2\, (*"z1 + z1 = 2 * z1"*) walther@60358: walther@60385: \<^rule_thm>\addiere_vor_minus\, (*"[| l is_num; m is_num |] ==> -(l * v) + m * v = (-l + m) *v"*) walther@60385: \<^rule_thm>\addiere_eins_vor_minus\, (*"[| m is_num |] ==> - v + m * v = (-1 + m) * v"*) walther@60385: \<^rule_thm>\subtrahiere_vor_minus\, (*"[| l is_num; m is_num |] ==> -(l * v) - m * v = (-l - m) *v"*) walther@60385: \<^rule_thm>\subtrahiere_eins_vor_minus\], (*"[| m is_num |] ==> - v - m * v = (-1 - m) * v"*) walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val verschoenere = walther@59851: Rule_Def.Repeat{id = "verschoenere", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: erls = Rule_Set.append_rules "erls_verschoenere" Rule_Set.empty walther@60358: [\<^rule_eval>\PolyMinus.kleiner\ (eval_kleiner "")], walther@60358: rules = [ walther@60358: \<^rule_thm>\vorzeichen_minus_weg1\, (*"l kleiner 0 ==> a + l * b = a - -l * b"*) walther@60358: \<^rule_thm>\vorzeichen_minus_weg2\, (*"l kleiner 0 ==> a - l * b = a + -l * b"*) walther@60358: \<^rule_thm>\vorzeichen_minus_weg3\, (*"l kleiner 0 ==> k + a - l * b = k + a + -l * b"*) walther@60358: \<^rule_thm>\vorzeichen_minus_weg4\, (*"l kleiner 0 ==> k - a - l * b = k - a + -l * b"*) neuper@37950: walther@60358: \<^rule_eval>\times\ (**)(eval_binop "#mult_"), neuper@37950: walther@60358: \<^rule_thm>\mult_zero_left\, (*"0 * z = 0"*) walther@60358: \<^rule_thm>\mult_1_left\, (*"1 * z = z"*) walther@60358: \<^rule_thm>\add_0_left\, (*"0 + z = z"*) walther@60358: \<^rule_thm>\null_minus\, (*"0 - a = -a"*) walther@60358: \<^rule_thm>\vor_minus_mal\], (*"- a * b = (-a) * b"*) walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val klammern_aufloesen = walther@59851: Rule_Def.Repeat{id = "klammern_aufloesen", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], erls = Rule_Set.Empty, walther@60358: rules = [ walther@60358: \<^rule_thm_sym>\add.assoc\, (*"a + (b + c) = (a + b) + c"*) walther@60358: \<^rule_thm>\klammer_plus_minus\, (*"a + (b - c) = (a + b) - c"*) walther@60358: \<^rule_thm>\klammer_minus_plus\, (*"a - (b + c) = (a - b) - c"*) walther@60358: \<^rule_thm>\klammer_minus_minus\], (*"a - (b - c) = (a - b) + c"*) walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val klammern_ausmultiplizieren = walther@59851: Rule_Def.Repeat{id = "klammern_ausmultiplizieren", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], erls = Rule_Set.Empty, walther@60358: rules = [ walther@60358: \<^rule_thm>\distrib_right\, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*) walther@60358: \<^rule_thm>\distrib_left\, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*) walther@60358: walther@60358: \<^rule_thm>\klammer_mult_minus\, (*"a * (b - c) = a * b - a * c"*) walther@60358: \<^rule_thm>\klammer_minus_mult\], (*"(b - c) * a = b * a - c * a"*) walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val ordne_monome = walther@60358: Rule_Def.Repeat { Walther@60509: id = "ordne_monome", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: erls = Rule_Set.append_rules "erls_ordne_monome" Rule_Set.empty [ walther@60358: \<^rule_eval>\PolyMinus.kleiner\ (eval_kleiner ""), walther@60358: \<^rule_eval>\Prog_Expr.is_atom\ (Prog_Expr.eval_is_atom "")], walther@60358: rules = [ walther@60358: \<^rule_thm>\tausche_mal\, (*"[| b is_atom; a kleiner b |] ==> (b * a) = (a * b)"*) walther@60358: \<^rule_thm>\tausche_vor_mal\, (*"[| b is_atom; a kleiner b |] ==> (-b * a) = (-a * b)"*) walther@60358: \<^rule_thm>\tausche_mal_mal\, (*"[| c is_atom; b kleiner c |] ==> (a * c * b) = (a * b *c)"*) walther@60358: \<^rule_thm>\x_quadrat\], (*"(x * a) * a = x * a \ 2"*) walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val rls_p_33 = walther@60358: Rule_Set.append_rules "rls_p_33" Rule_Set.empty [ walther@60358: Rule.Rls_ ordne_alphabetisch, wneuper@59416: Rule.Rls_ fasse_zusammen, walther@60358: Rule.Rls_ verschoenere]; neuper@37950: val rls_p_34 = walther@60358: Rule_Set.append_rules "rls_p_34" Rule_Set.empty [ walther@60358: Rule.Rls_ klammern_aufloesen, walther@60358: Rule.Rls_ ordne_alphabetisch, walther@60358: Rule.Rls_ fasse_zusammen, walther@60358: Rule.Rls_ verschoenere]; neuper@37950: val rechnen = walther@60358: Rule_Set.append_rules "rechnen" Rule_Set.empty [ walther@60358: \<^rule_eval>\times\ (**)(eval_binop "#mult_"), walther@60358: \<^rule_eval>\plus\ (**)(eval_binop "#add_"), walther@60358: \<^rule_eval>\minus\ (**)(eval_binop "#subtr_")]; wneuper@59472: \ wenzelm@60289: rule_set_knowledge wenzelm@60286: ordne_alphabetisch = \prep_rls' ordne_alphabetisch\ and wenzelm@60286: fasse_zusammen = \prep_rls' fasse_zusammen\ and wenzelm@60286: verschoenere = \prep_rls' verschoenere\ and wenzelm@60286: ordne_monome = \prep_rls' ordne_monome\ and wenzelm@60286: klammern_aufloesen = \prep_rls' klammern_aufloesen\ and wenzelm@60286: klammern_ausmultiplizieren = \prep_rls' klammern_ausmultiplizieren\ neuper@37950: neuper@37950: (** problems **) wenzelm@60306: wenzelm@60306: problem pbl_vereinf_poly : "polynom/vereinfachen" = \Rule_Set.Empty\ wenzelm@60306: wenzelm@60306: problem pbl_vereinf_poly_minus : "plus_minus/polynom/vereinfachen" = walther@60358: \Rule_Set.append_rules "prls_pbl_vereinf_poly" Rule_Set.empty [ walther@60358: \<^rule_eval>\is_polyexp\ (eval_is_polyexp ""), wenzelm@60306: \<^rule_eval>\Prog_Expr.matchsub\ (Prog_Expr.eval_matchsub ""), walther@60358: \<^rule_thm>\or_true\, (*"(?a | True) = True"*) walther@60358: \<^rule_thm>\or_false\, (*"(?a | False) = ?a"*) walther@60358: \<^rule_thm>\not_true\, (*"(~ True) = False"*) walther@60358: \<^rule_thm>\not_false\ (*"(~ False) = True"*)]\ Walther@60449: Method_Ref: "simplification/for_polynomials/with_minus" wenzelm@60306: CAS: "Vereinfache t_t" wenzelm@60306: Given: "Term t_t" wenzelm@60306: Where: wenzelm@60306: "t_t is_polyexp" wenzelm@60306: "Not (matchsub (?a + (?b + ?c)) t_t | wenzelm@60306: matchsub (?a + (?b - ?c)) t_t | wenzelm@60306: matchsub (?a - (?b + ?c)) t_t | wenzelm@60306: matchsub (?a + (?b - ?c)) t_t )" wenzelm@60306: "Not (matchsub (?a * (?b + ?c)) t_t | wenzelm@60306: matchsub (?a * (?b - ?c)) t_t | wenzelm@60306: matchsub ((?b + ?c) * ?a) t_t | wenzelm@60306: matchsub ((?b - ?c) * ?a) t_t )" wenzelm@60306: Find: "normalform n_n" wenzelm@60306: wenzelm@60306: problem pbl_vereinf_poly_klammer : "klammer/polynom/vereinfachen" = walther@60358: \Rule_Set.append_rules "prls_pbl_vereinf_poly_klammer" Rule_Set.empty [ walther@60358: \<^rule_eval>\is_polyexp\ (eval_is_polyexp ""), wenzelm@60306: \<^rule_eval>\Prog_Expr.matchsub\ (Prog_Expr.eval_matchsub ""), walther@60358: \<^rule_thm>\or_true\, (*"(?a | True) = True"*) walther@60358: \<^rule_thm>\or_false\, (*"(?a | False) = ?a"*) walther@60358: \<^rule_thm>\not_true\, (*"(~ True) = False"*) walther@60358: \<^rule_thm>\not_false\ (*"(~ False) = True"*)]\ Walther@60449: Method_Ref: "simplification/for_polynomials/with_parentheses" wenzelm@60306: CAS: "Vereinfache t_t" wenzelm@60306: Given: "Term t_t" wenzelm@60306: Where: wenzelm@60306: "t_t is_polyexp" wenzelm@60306: "Not (matchsub (?a * (?b + ?c)) t_t | wenzelm@60306: matchsub (?a * (?b - ?c)) t_t | wenzelm@60306: matchsub ((?b + ?c) * ?a) t_t | wenzelm@60306: matchsub ((?b - ?c) * ?a) t_t )" wenzelm@60306: Find: "normalform n_n" wenzelm@60306: wenzelm@60306: problem pbl_vereinf_poly_klammer_mal : "binom_klammer/polynom/vereinfachen" = wenzelm@60306: \Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*) wenzelm@60306: \<^rule_eval>\is_polyexp\ (eval_is_polyexp "")]\ Walther@60449: Method_Ref: "simplification/for_polynomials/with_parentheses_mult" wenzelm@60306: CAS: "Vereinfache t_t" wenzelm@60306: Given: "Term t_t" wenzelm@60306: Where: "t_t is_polyexp" wenzelm@60306: Find: "normalform n_n" wenzelm@60306: wenzelm@60306: problem pbl_probe : "probe" = \Rule_Set.Empty\ wenzelm@60306: wenzelm@60306: problem pbl_probe_poly : "polynom/probe" = wenzelm@60306: \Rule_Set.append_rules "prls_pbl_probe_poly" Rule_Set.empty [(*for preds in where_*) wenzelm@60306: \<^rule_eval>\is_polyexp\ (eval_is_polyexp "")]\ Walther@60449: Method_Ref: "probe/fuer_polynom" wenzelm@60306: CAS: "Probe e_e w_w" wenzelm@60306: Given: "Pruefe e_e" "mitWert w_w" wenzelm@60306: Where: "e_e is_polyexp" wenzelm@60306: Find: "Geprueft p_p" wenzelm@60306: wenzelm@60306: problem pbl_probe_bruch : "bruch/probe" = wenzelm@60306: \Rule_Set.append_rules "prls_pbl_probe_bruch" Rule_Set.empty [(*for preds in where_*) wenzelm@60306: \<^rule_eval>\is_ratpolyexp\ (eval_is_ratpolyexp "")]\ Walther@60449: Method_Ref: "probe/fuer_bruch" wenzelm@60306: CAS: "Probe e_e w_w" wenzelm@60306: Given: "Pruefe e_e" "mitWert w_w" wenzelm@60306: Where: "e_e is_ratpolyexp" wenzelm@60306: Find: "Geprueft p_p" neuper@37950: neuper@37950: (** methods **) wneuper@59545: wneuper@59504: partial_function (tailrec) simplify :: "real \ real" wneuper@59504: where walther@59635: "simplify t_t = ( walther@59635: (Repeat( walther@59637: (Try (Rewrite_Set ''ordne_alphabetisch'')) #> walther@59637: (Try (Rewrite_Set ''fasse_zusammen'')) #> walther@59635: (Try (Rewrite_Set ''verschoenere''))) walther@59635: ) t_t)" wenzelm@60303: wenzelm@60303: method met_simp_poly_minus : "simplification/for_polynomials/with_minus" = wenzelm@60303: \{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, walther@60358: prls = Rule_Set.append_rules "prls_met_simp_poly_minus" Rule_Set.empty [ walther@60358: \<^rule_eval>\is_polyexp\ (eval_is_polyexp ""), walther@60358: \<^rule_eval>\Prog_Expr.matchsub\ (Prog_Expr.eval_matchsub ""), walther@60358: \<^rule_thm>\and_true\, (*"(?a & True) = ?a"*) walther@60358: \<^rule_thm>\and_false\, (*"(?a & False) = False"*) walther@60358: \<^rule_thm>\not_true\, (*"(~ True) = False"*) walther@60358: \<^rule_thm>\not_false\ (*"(~ False) = True"*)], wenzelm@60303: crls = Rule_Set.empty, errpats = [], nrls = rls_p_33}\ wenzelm@60303: Program: simplify.simps wenzelm@60303: Given: "Term t_t" wenzelm@60303: Where: wenzelm@60303: "t_t is_polyexp" wenzelm@60303: "Not (matchsub (?a + (?b + ?c)) t_t | wenzelm@60303: matchsub (?a + (?b - ?c)) t_t | wenzelm@60303: matchsub (?a - (?b + ?c)) t_t | wenzelm@60303: matchsub (?a + (?b - ?c)) t_t)" wenzelm@60303: Find: "normalform n_n" wneuper@59545: wneuper@59504: partial_function (tailrec) simplify2 :: "real \ real" wneuper@59504: where walther@59635: "simplify2 t_t = ( walther@59635: (Repeat( walther@59637: (Try (Rewrite_Set ''klammern_aufloesen'')) #> walther@59637: (Try (Rewrite_Set ''ordne_alphabetisch'')) #> walther@59637: (Try (Rewrite_Set ''fasse_zusammen'')) #> walther@59635: (Try (Rewrite_Set ''verschoenere''))) walther@59635: ) t_t)" wenzelm@60303: wenzelm@60303: method met_simp_poly_parenth : "simplification/for_polynomials/with_parentheses" = wenzelm@60303: \{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, wenzelm@60303: prls = Rule_Set.append_rules "simplification_for_polynomials_prls" Rule_Set.empty wenzelm@60303: [(*for preds in where_*) \<^rule_eval>\is_polyexp\ (eval_is_polyexp"")], wenzelm@60303: crls = Rule_Set.empty, errpats = [], nrls = rls_p_34}\ wenzelm@60303: Program: simplify2.simps wenzelm@60303: Given: "Term t_t" wenzelm@60303: Where: "t_t is_polyexp" wenzelm@60303: Find: "normalform n_n" wneuper@59545: wneuper@59504: partial_function (tailrec) simplify3 :: "real \ real" wneuper@59504: where walther@59635: "simplify3 t_t = ( walther@59635: (Repeat( walther@59637: (Try (Rewrite_Set ''klammern_ausmultiplizieren'')) #> walther@59637: (Try (Rewrite_Set ''discard_parentheses'')) #> walther@59637: (Try (Rewrite_Set ''ordne_monome'')) #> walther@59637: (Try (Rewrite_Set ''klammern_aufloesen'')) #> walther@59637: (Try (Rewrite_Set ''ordne_alphabetisch'')) #> walther@59637: (Try (Rewrite_Set ''fasse_zusammen'')) #> walther@59635: (Try (Rewrite_Set ''verschoenere''))) walther@59635: ) t_t)" wenzelm@60303: wenzelm@60303: method met_simp_poly_parenth_mult : "simplification/for_polynomials/with_parentheses_mult" = wenzelm@60303: \{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, wenzelm@60303: prls = Rule_Set.append_rules "simplification_for_polynomials_prls" Rule_Set.empty wenzelm@60303: [(*for preds in where_*) \<^rule_eval>\is_polyexp\ (eval_is_polyexp"")], wenzelm@60303: crls = Rule_Set.empty, errpats = [], nrls = rls_p_34}\ wenzelm@60303: Program: simplify3.simps wenzelm@60303: Given: "Term t_t" wenzelm@60303: Where: "t_t is_polyexp" wenzelm@60303: Find: "normalform n_n" wenzelm@60303: wenzelm@60303: method met_probe : "probe" = wenzelm@60303: \{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.Empty, crls = Rule_Set.empty, wenzelm@60303: errpats = [], nrls = Rule_Set.Empty}\ wneuper@59545: wneuper@59504: partial_function (tailrec) mache_probe :: "bool \ bool list \ bool" wneuper@59504: where walther@59635: "mache_probe e_e w_w = ( walther@59635: let walther@59635: e_e = Take e_e; walther@59635: e_e = Substitute w_w e_e walther@59635: in ( walther@59635: Repeat ( walther@59637: (Try (Repeat (Calculate ''TIMES''))) #> walther@59637: (Try (Repeat (Calculate ''PLUS'' ))) #> walther@59635: (Try (Repeat (Calculate ''MINUS'')))) walther@59635: ) e_e)" wenzelm@60303: wenzelm@60303: method met_probe_poly : "probe/fuer_polynom" = wenzelm@60303: \{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, wenzelm@60303: prls = Rule_Set.append_rules "prls_met_probe_bruch" Rule_Set.empty wenzelm@60303: [(*for preds in where_*) \<^rule_eval>\is_ratpolyexp\ (eval_is_ratpolyexp "")], wenzelm@60303: crls = Rule_Set.empty, errpats = [], nrls = rechnen}\ wenzelm@60303: Program: mache_probe.simps wenzelm@60303: Given: "Pruefe e_e" "mitWert w_w" wenzelm@60303: Where: "e_e is_polyexp" wenzelm@60303: Find: "Geprueft p_p" wenzelm@60303: wenzelm@60303: method met_probe_bruch : "probe/fuer_bruch" = wenzelm@60303: \{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, wenzelm@60303: prls = Rule_Set.append_rules "prls_met_probe_bruch" Rule_Set.empty wenzelm@60303: [(*for preds in where_*) \<^rule_eval>\is_ratpolyexp\ (eval_is_ratpolyexp "")], wenzelm@60303: crls = Rule_Set.empty, errpats = [], nrls = Rule_Set.Empty}\ wenzelm@60303: Given: "Pruefe e_e" "mitWert w_w" wenzelm@60303: Where: "e_e is_ratpolyexp" wenzelm@60303: Find: "Geprueft p_p" wenzelm@60303: wenzelm@60303: ML \ walther@60278: \ ML \ wneuper@59472: \ neuper@37906: neuper@37906: end neuper@37906: