test/Tools/isac/ADDTESTS/course/phst11/T2_Rewriting.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Sat, 24 Feb 2018 16:09:24 +0100
changeset 59382 364ce4699452
parent 59348 ddfabb53082c
child 59395 862eb17f9e16
permissions -rw-r--r--
Rewrite is broken; intermediate state for repair
     1 (*
     2 cd /usr/local/Isabelle/test/Tools/isac/ADDTESTS/course/
     3 /usr/local/Isabelle/bin/isabelle jedit -l Isac T2_Rewriting.thy &
     4 *)
     5 
     6 theory T2_Rewriting imports Isac begin
     7 
     8 chapter {* Rewriting *}
     9 
    10 text {* \emph{Rewriting} is a technique of Symbolic Computation, which is 
    11   appropriate to make a 'transparent system', because it is intuitively 
    12   comprehensible. For a thourogh introduction see the book of Tobias Nipkow, 
    13   http://www4.informatik.tu-muenchen.de/~nipkow/TRaAT/
    14 
    15 section {* Introduction to rewriting *}
    16 
    17 text {* Rewriting creates calculations which look like written by hand; see the 
    18   ISAC tutoring system ! ISAC finds the rules automatically. Here we start by 
    19   telling the rules ourselves.
    20   Let's differentiate after we have identified the rules for differentiation, as 
    21   found in ~~/src/Tools/isac/Knowledge/Diff.thy:
    22 *}
    23 ML {*
    24 val diff_sum = num_str @{thm diff_sum};
    25 val diff_pow = num_str @{thm diff_pow};
    26 val diff_var = num_str @{thm diff_var};
    27 val diff_const = num_str @{thm diff_const};
    28 *}
    29 text {* Looking at the rules (abbreviated by 'thm' above), we see the 
    30   differential operator abbreviated by 'd_d ?bdv', where '?bdv' is the bound 
    31   variable.
    32   Can you read diff_const in the Ouput window ?
    33 
    34   Please, skip this introductory ML-section in the first go ...*}
    35 ML {*
    36 (*default_print_depth 1;*)
    37 val (thy, ro, er, inst) = 
    38     (@{theory "Isac"}, tless_true, eval_rls, [(@{term "bdv::real"}, @{term "x::real"})]);
    39 *}
    40 text {* ... and  let us differentiate the term t: *}
    41 ML {*
    42 val t = (Thm.term_of o the o (parse thy)) "d_d x (x^^^2 + x + y)";
    43 
    44 val SOME (t, _) = Rewrite.rewrite_inst_ thy ro er true inst diff_sum t; term2str t;
    45 val SOME (t, _) = Rewrite.rewrite_inst_ thy ro er true inst diff_sum t; term2str t;
    46 val SOME (t, _) = Rewrite.rewrite_inst_ thy ro er true inst diff_pow t; term2str t;
    47 val SOME (t, _) = Rewrite.rewrite_inst_ thy ro er true inst diff_var t; term2str t;
    48 val SOME (t, _) = Rewrite.rewrite_inst_ thy ro er true inst diff_const t; term2str t;
    49 *}
    50 text {* Please, scoll up the Output-window to check the 5 steps of rewriting !
    51   You might not be satisfied by the result "2 * x ^^^ (2 - 1) + 1 + 0".
    52 
    53   ISAC has a set of rules called 'make_polynomial', which simplifies the result:
    54 *}
    55 ML {* 
    56 val SOME (t, _) = Rewrite.rewrite_set_ thy true make_polynomial t; term2str t;
    57 trace_rewrite := false;
    58 *}
    59 
    60 section {* Note on bound variables *}
    61 text {* You may have noticed that rewrite_ above could distinguish between x 
    62   (d_d x x = 1) and y (d_d x y = 0). ISAC does this by instantiating theorems
    63   before application: given [(@{term "bdv::real"}, @{term "x::real"})] the 
    64   theorem diff_sum becomes "d_d x (?u + ?v) = d_d x ?u + d_d x ?v".
    65 
    66   Isabelle does this differently by variables bound by a 'lambda' %, see
    67   http://isabelle.in.tum.de/dist/library/HOL/HOL-Multivariate_Analysis/Derivative.html
    68 *}
    69 ML {*
    70 val t = @{term "%x. x^2 + x + y"}; atomwy t; term2str t;
    71 *}
    72 text {* Since this notation does not conform to present high-school matheatics
    73   ISAC introduced the 'bdv' mechanism. This mechanism is also used for equation
    74   solving in ISAC.
    75 *}
    76 
    77 section {* Conditional and ordered rewriting *}
    78 text {* We have already seen conditional rewriting above when we used the rule
    79   diff_const; let us try: *}
    80 ML {*
    81 val t1 = (Thm.term_of o the o (parse thy)) "d_d x (a*BC*x*z)";
    82 Rewrite.rewrite_inst_ thy ro er true inst diff_const t1;
    83 
    84 val t2 = (Thm.term_of o the o (parse thy)) "d_d x (a*BC*y*z)";
    85 Rewrite.rewrite_inst_ thy ro er true inst diff_const t2;
    86 *}
    87 text {* For term t1 the assumption 'not (x occurs_in "a*BC*x*z")' is false, 
    88   since x occurs in t1 actually; thus the rule following implication '==>' is 
    89   not applied and rewrite_inst_ returns NONE.
    90   For term t2 is 'not (x occurs_in "a*BC*y*z")' true, thus the rule is applied.
    91 *}
    92 
    93 subsection {* ordered rewriting *}
    94 text {* Let us start with an example; in order to see what is going on we tell
    95   Isabelle to show all brackets:
    96 *}
    97 ML {*
    98 (*show_brackets := true; TODO*)
    99 val t0 = (Thm.term_of o the o (parse thy)) "2*a + 3*b + 4*a"; term2str t0;
   100 (*show_brackets := false;*)
   101 *}
   102 text {* Now we want to bring 4*a close to 2*a in order to get 6*a:
   103 *}
   104 ML {*
   105 val SOME (t, _) = Rewrite.rewrite_ thy ro er true @{thm add.assoc} t0; term2str t;
   106 val SOME (t, _) = Rewrite.rewrite_ thy ro er true @{thm add.left_commute} t; term2str t;
   107 val SOME (t, _) = Rewrite.rewrite_ thy ro er true @{thm add.commute} t; term2str t;
   108 val SOME (t, _) = Rewrite.rewrite_ thy ro er true @{thm real_num_collect} t; term2str t;
   109 *}
   110 text {* That is fine, we just need to add 2+4 !!!!! See the next section below.
   111 
   112   But we cannot automate such ordering with what we know so far: If we put
   113   add.assoc, add.left_commute and add.commute into one ruleset (your have used
   114   ruleset 'make_polynomial' already), then all the rules are applied as long
   115   as one rule is applicable (that is the way such rulesets work).
   116   Try to step through the ML-sections without skipping one of them ...
   117 *}
   118 ML {*val SOME (t, _) = Rewrite.rewrite_ thy ro er true @{thm add.commute} t; term2str t*}
   119 ML {*val SOME (t, _) = Rewrite.rewrite_ thy ro er true @{thm add.commute} t; term2str t*}
   120 ML {*val SOME (t, _) = Rewrite.rewrite_ thy ro er true @{thm add.commute} t; term2str t*}
   121 ML {*val SOME (t, _) = Rewrite.rewrite_ thy ro er true @{thm add.commute} t; term2str t*}
   122 text {* ... you can go forever, the ruleset is 'not terminating'.
   123   The theory of rewriting makes this kind of rulesets terminate by the use of
   124   'rewrite orders': 
   125   Given two terms t1 and t2 we describe rewriting by: t1 ~~> t2. Then
   126   'ordered rewriting' is: t2 < t1 ==> t1 ~~> t2. That means, a rule is only
   127   applied, when the result t2 is 'smaller', '<', than the one to be rewritten.
   128   Defining such a '<' is not trivial, see ~~/src/Tools/isacKnowledge/Poly.thy
   129   for 'fun has_degree_in' etc.
   130 *}
   131 
   132 subsection {* Simplification in ISAC *}
   133 text {* 
   134   With the introduction into rewriting, ordered rewriting and conditional
   135   rewriting we have seen all what is necessary for the practice of rewriting.
   136 
   137   One basic technique of 'symbolic computation' which uses rewriting is
   138   simplification, that means: transform terms into an equivalent form which is
   139   as simple as possible.
   140   Isabelle has powerful and efficient simplifiers. Nevertheless, ISAC has another
   141   kind of simplifiers, which groups rulesets such that the trace of rewrites is 
   142   more readable in ISAC's worksheets.
   143 
   144   Here are examples of of how ISAC's simplifier work:
   145 *}
   146 ML {*
   147 (*show_brackets := false; TODO*)
   148 val t1 = (Thm.term_of o the o (parse thy)) "(a - b) * (a^^^2 + a*b + b^^^2)";
   149 val SOME (t, _) = Rewrite.rewrite_set_ thy true make_polynomial t1; term2str t;
   150 *}
   151 ML {*
   152 val t2 = (Thm.term_of o the o (parse thy)) 
   153   "(2 / (x + 3) + 2 / (x - 3)) / (8 * x / (x ^^^ 2 - 9))";
   154 val SOME (t, _) = Rewrite.rewrite_set_ thy true norm_Rational t2; term2str t;
   155 *}
   156 text {* The simplifiers are quite busy when finding the above results. you can
   157   watch them at work by setting the switch 'trace_rewrite:*}
   158 ML {*
   159 trace_rewrite := false;
   160 tracing "+++begin++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++";
   161 val SOME (t, _) = Rewrite.rewrite_set_ thy true norm_Rational t2; term2str t;
   162 tracing "+++end++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++";
   163 trace_rewrite := false;
   164 *}
   165 text {* You see what happend when you click the checkbox <Tracing> on the bar
   166   separating this window from the Output-window.
   167 
   168   So, it might be better to take simpler examples for watching the simplifiers.
   169 *}
   170 
   171 
   172 section {* Experiments with a simplifier conserving minus *}
   173 
   174 text {* We conclude the section on rewriting with starting into an experimental
   175   development. This development has been urged by teachers using ISAC for
   176   introduction to algebra with students at the age of 12-14.
   177 
   178   The teachers requested ISAC to keep the minus, for instance in the above 
   179   result "a^3 + -1 * b^3" (here ISAC should write "a^3  - * b^3") and also
   180   in all intermediate steps.
   181 
   182   So we started to develop (in German !) such a simplifier in 
   183   ~~/src/Tools/isac/Knowledge/PolyMinus.thy
   184 *}
   185 
   186 subsection {* What already works *}
   187 ML {*
   188 val t2 = (Thm.term_of o the o (parse thy)) 
   189   "5*e + 6*f - 8*g - 9 - 7*e - 4*f + 10*g + 12";
   190 val SOME (t, _) = Rewrite.rewrite_set_ thy true rls_p_33 t2; term2str t;
   191 *}
   192 text {* Try your own examples ! *}
   193 
   194 subsection {* This raises questions about didactics *}
   195 text {* Oberserve the '-' ! That works out. But see the efforts in PolyMinus.thy
   196 *}
   197 ML {*
   198 @{thm subtrahiere};
   199 @{thm subtrahiere_von_1};
   200 @{thm subtrahiere_1};
   201 *}
   202 text {* That would not be so bad. But it is only a little part of what else is
   203   required:
   204 *}
   205 ML {*
   206 @{thm subtrahiere_x_plus_minus};
   207 @{thm subtrahiere_x_plus1_minus};
   208 @{thm subtrahiere_x_plus_minus1};
   209 @{thm subtrahiere_x_minus_plus};
   210 @{thm subtrahiere_x_minus1_plus};
   211 @{thm subtrahiere_x_minus_plus1};
   212 @{thm subtrahiere_x_minus_minus};
   213 @{thm subtrahiere_x_minus1_minus};
   214 @{thm subtrahiere_x_minus_minus1};
   215 *}
   216 text {* So, learning so many rules, and learning to apply these rules, is futile.
   217   Actually, most of our students are unable to apply theorems.
   218 
   219   But if you look at 'make_polynomial' or even 'norm_Rational' you see, 
   220   that these general simplifiers require about 10% than rulesets for '-'.
   221 
   222   So, we might have better chances to teach our student to apply theorems
   223   without '-' ?
   224 *}
   225 
   226 subsection {* This does not yet work *}
   227 ML {*
   228 val t = (Thm.term_of o the o (parse thy)) 
   229   "(2*a - 5*b) * (2*a + 5*b)";
   230 Rewrite.rewrite_set_ thy true rls_p_33 t; term2str t;
   231 *}
   232 
   233 end