test/Tools/isac/ADDTESTS/course/phst11/T2_Rewriting.thy
author wneuper <Walther.Neuper@jku.at>
Mon, 01 Jan 2024 11:31:16 +0100
changeset 60789 8fa678b678e8
parent 60675 d841c720d288
permissions -rw-r--r--
Doc/Specify_Phase 4: start use antiquotations from isar-ref
     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 $ISABELLE_ISAC/Knowledge/Diff.thy:
    23 \<close>
    24 ML \<open>
    25 val diff_sum = @{thm diff_sum};
    26 val diff_pow = @{thm diff_pow};
    27 val diff_var = @{thm diff_var};
    28 val diff_const = @{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 val ctxt = Proof_Context.init_global thy; (*required for parsing*)
    41 \<close>
    42 text \<open>... and  let us differentiate the term t:\<close>
    43 ML \<open>
    44 val t = ParseC.parse_test ctxt "d_d x (x\<up>2 + x + y)";
    45 
    46 val SOME (t, _) = Rewrite.rewrite_inst_ ctxt ro er true inst diff_sum t; UnparseC.term @{context} t;
    47 val SOME (t, _) = Rewrite.rewrite_inst_ ctxt ro er true inst diff_sum t; UnparseC.term @{context} t;
    48 val SOME (t, _) = Rewrite.rewrite_inst_ ctxt ro er true inst diff_pow t; UnparseC.term @{context} t;
    49 val SOME (t, _) = Rewrite.rewrite_inst_ ctxt ro er true inst diff_var t; UnparseC.term @{context} t;
    50 val SOME (t, _) = Rewrite.rewrite_inst_ ctxt ro er true inst diff_const t; UnparseC.term @{context} t;
    51 \<close>
    52 text \<open>Please, scoll up the Output-window to check the 5 steps of rewriting !
    53   You might not be satisfied by the result "2 * x \<up> (2 - 1) + 1 + 0".
    54 
    55   ISAC has a set of rules called 'make_polynomial', which simplifies the result:
    56 \<close>
    57 ML \<open>
    58 val SOME (t, _) = Rewrite.rewrite_set_ ctxt true make_polynomial t; UnparseC.term @{context} t;
    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"};
    72 TermC.atom_write_detail @{context} t; UnparseC.term @{context} t;
    73 \<close>
    74 text \<open>Since this notation does not conform to present high-school matheatics
    75   ISAC introduced the 'bdv' mechanism. This mechanism is also used for equation
    76   solving in ISAC.
    77 \<close>
    78 
    79 section \<open>Conditional and ordered rewriting\<close>
    80 text \<open>We have already seen conditional rewriting above when we used the rule
    81   diff_const; let us try:\<close>
    82 ML \<open>
    83 val t1 = ParseC.parse_test ctxt "d_d x (a*BC*x*z)";
    84 Rewrite.rewrite_inst_ ctxt ro er true inst diff_const t1;
    85 
    86 val t2 = ParseC.parse_test ctxt "d_d x (a*BC*y*z)";
    87 Rewrite.rewrite_inst_ ctxt ro er true inst diff_const t2;
    88 \<close>
    89 text \<open>For term t1 the assumption 'not (x occurs_in "a*BC*x*z")' is false, 
    90   since x occurs in t1 actually; thus the rule following implication '==>' is 
    91   not applied and rewrite_inst_ returns NONE.
    92   For term t2 is 'not (x occurs_in "a*BC*y*z")' true, thus the rule is applied.
    93 \<close>
    94 
    95 subsection \<open>ordered rewriting\<close>
    96 text \<open>Let us start with an example; in order to see what is going on we tell
    97   Isabelle to show all brackets:
    98 \<close>
    99 ML \<open>
   100 (*show_brackets := true; TODO*)
   101 val t0 = ParseC.parse_test ctxt "2*a + 3*b + 4*a::real"; UnparseC.term @{context} t0;
   102 (*show_brackets := false;*)
   103 \<close>
   104 text \<open>Now we want to bring 4*a close to 2*a in order to get 6*a:
   105 \<close>
   106 ML \<open>
   107 val SOME (t, _) = Rewrite.rewrite_ ctxt ro er true @{thm add.assoc} t0; UnparseC.term @{context} t;
   108 val SOME (t, _) = Rewrite.rewrite_ ctxt ro er true @{thm add.left_commute} t; UnparseC.term @{context} t;
   109 val SOME (t, _) = Rewrite.rewrite_ ctxt ro er true @{thm add.commute} t; UnparseC.term @{context} t;
   110 val SOME (t, _) = Rewrite.rewrite_ ctxt ro er true @{thm real_num_collect} t; UnparseC.term @{context} t;
   111 \<close>
   112 text \<open>That is fine, we just need to add 2+4 !!!!! See the next section below.
   113 
   114   But we cannot automate such ordering with what we know so far: If we put
   115   add.assoc, add.left_commute and add.commute into one ruleset (your have used
   116   ruleset 'make_polynomial' already), then all the rules are applied as long
   117   as one rule is applicable (that is the way such rulesets work).
   118   Try to step through the ML-sections without skipping one of them ...
   119 \<close>
   120 ML \<open>val SOME (t, _) = Rewrite.rewrite_ ctxt ro er true @{thm add.commute} t; UnparseC.term @{context} t\<close>
   121 ML \<open>val SOME (t, _) = Rewrite.rewrite_ ctxt ro er true @{thm add.commute} t; UnparseC.term @{context} t\<close>
   122 ML \<open>val SOME (t, _) = Rewrite.rewrite_ ctxt ro er true @{thm add.commute} t; UnparseC.term @{context} t\<close>
   123 ML \<open>val SOME (t, _) = Rewrite.rewrite_ ctxt ro er true @{thm add.commute} t; UnparseC.term @{context} t\<close>
   124 text \<open>... you can go forever, the ruleset is 'not terminating'.
   125   The theory of rewriting makes this kind of rulesets terminate by the use of
   126   'rewrite orders': 
   127   Given two terms t1 and t2 we describe rewriting by: t1 ~~> t2. Then
   128   'ordered rewriting' is: t2 < t1 ==> t1 ~~> t2. That means, a rule is only
   129   applied, when the result t2 is 'smaller', '<', than the one to be rewritten.
   130   Defining such a '<' is not trivial, see ~~/src/Tools/isacKnowledge/Poly.thy
   131   for 'fun has_degree_in' etc.
   132 \<close>
   133 
   134 subsection \<open>Simplification in ISAC\<close>
   135 text \<open>
   136   With the introduction into rewriting, ordered rewriting and conditional
   137   rewriting we have seen all what is necessary for the practice of rewriting.
   138 
   139   One basic technique of 'symbolic computation' which uses rewriting is
   140   simplification, that means: transform terms into an equivalent form which is
   141   as simple as possible.
   142   Isabelle has powerful and efficient simplifiers. Nevertheless, ISAC has another
   143   kind of simplifiers, which groups rulesets such that the trace of rewrites is 
   144   more readable in ISAC's worksheets.
   145 
   146   Here are examples of of how ISAC's simplifier work:
   147 \<close>
   148 ML \<open>
   149 (*show_brackets := false; TODO*)
   150 val t1 = ParseC.parse_test ctxt "(a - b) * (a\<up>2 + a*b + b\<up>2)";
   151 val SOME (t, _) = Rewrite.rewrite_set_ ctxt true make_polynomial t1; UnparseC.term @{context} t;
   152 \<close>
   153 ML \<open>
   154 val t2 = ParseC.parse_test ctxt 
   155   "(2 / (x + 3) + 2 / (x - 3)) / (8 * x / (x \<up> 2 - 9))";
   156 val SOME (t, _) = Rewrite.rewrite_set_ ctxt true norm_Rational t2; UnparseC.term @{context} t;
   157 \<close>
   158 text \<open>The simplifiers are quite busy when finding the above results. you can
   159   watch them at work by setting the switch 'Rewrite.trace_on:\<close>
   160 ML \<open>
   161 tracing "+++begin++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++";
   162 val SOME (t, _) = Rewrite.rewrite_set_ ctxt true norm_Rational t2; UnparseC.term @{context} t;
   163 tracing "+++end++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++";
   164 \<close>
   165 text \<open>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 \<close>
   170 
   171 
   172 section \<open>Experiments with a simplifier conserving minus\<close>
   173 
   174 text \<open>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   $ISABELLE_ISAC/Knowledge/PolyMinus.thy
   184 \<close>
   185 
   186 subsection \<open>What already works\<close>
   187 ML \<open>
   188 val t2 = ParseC.parse_test ctxt 
   189   "5*e + 6*f - 8*g - 9 - 7*e - 4*f + 10*g + 12::real";
   190 val SOME (t, _) = Rewrite.rewrite_set_ ctxt true rls_p_33 t2; UnparseC.term @{context} t;
   191 \<close>
   192 text \<open>Try your own examples !\<close>
   193 
   194 subsection \<open>This raises questions about didactics\<close>
   195 text \<open>Oberserve the '-' ! That works out. But see the efforts in PolyMinus.thy
   196 \<close>
   197 ML \<open>
   198 @{thm subtrahiere};
   199 @{thm subtrahiere_von_1};
   200 @{thm subtrahiere_1};
   201 \<close>
   202 text \<open>That would not be so bad. But it is only a little part of what else is
   203   required:
   204 \<close>
   205 ML \<open>
   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 \<close>
   216 text \<open>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 \<close>
   225 
   226 subsection \<open>This does not yet work\<close>
   227 ML \<open>
   228 val t = ParseC.parse_test ctxt 
   229   "(2*a - 5*b) * (2*a + 5*b)";
   230 Rewrite.rewrite_set_ ctxt true rls_p_33 t; UnparseC.term @{context} t;
   231 \<close>
   232 
   233 end