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