test/Tools/isac/ADDTESTS/course/CADGME/example_2.thy
author wneuper <Walther.Neuper@jku.at>
Sun, 09 Oct 2022 07:44:22 +0200
changeset 60565 f92963a33fe3
parent 60242 73ee61385493
child 60665 fad0cbfb586d
permissions -rw-r--r--
eliminate term2str in test/*
     1 
     2 theory example_2 imports Isac.Isac_Knowledge
     3 begin
     4 
     5 section\<open>Symbol Representation\<close>
     6 
     7 ML \<open>
     8   (* 
     9    * alpha -->  "</alpha>" 
    10    *)
    11   @{term "\<alpha> "};
    12   @{term "\<delta> "};
    13   @{term "\<phi> "};
    14   @{term "\<rho> "};
    15   UnparseC.term @{term "\<rho> "};
    16 \<close>
    17 
    18 section\<open>Working with Rulesets\<close>
    19 
    20 axiomatization where 
    21   rule1: "1 = \<delta>[n]" and
    22   rule2: "|| z || > 1 ==> z / (z - 1) = u [n]" and
    23   rule3: "|| z || < 1 ==> z / (z - 1) = -u [-n - 1]" and 
    24   rule4: "|| z || > || \<alpha> || ==> z / (z - \<alpha>) = \<alpha>\<up>n * u [n]"
    25 
    26 ML \<open>
    27   @{thm rule1};
    28   @{thm rule2};
    29   @{thm rule3};
    30   @{thm rule4};
    31 \<close>
    32 
    33 text\<open>Create Ruleset and insert rules\<close>
    34 
    35 ML \<open>
    36   val transform = Rule_Set.append_rules "transform" Rule_Set.empty
    37     [ Thm  ("rule3",ThmC.numerals_to_Free @{thm rule3}),
    38       Thm  ("rule4",ThmC.numerals_to_Free @{thm rule4}),
    39       Thm  ("rule1",ThmC.numerals_to_Free @{thm rule1})   
    40     ];
    41 \<close>
    42 
    43 text\<open>Apply the Ruleset to a term\<close>
    44 
    45 ML \<open>
    46   val t = parse_test @{context} "z / (z - 1) + z / (z - \<alpha>) + 1::real";
    47   val SOME (t', asm) = rewrite_set_ thy true transform t;
    48 \<close>
    49 
    50 ML\<open>
    51  UnparseC.term t';
    52  UnparseC.term t' = "z / (z - ?\<delta> [?n]) + z / (z - \<alpha>) + ?\<delta> [?n]";
    53  UnparseC.term t' = "?u [?n] + \<alpha>\<up>?n u[n] + ?\<delta> [?n]";
    54 \<close>
    55 
    56 text\<open>Problems: Rule~1 is applied before the expression is checked for Rule~4
    57       which would be the correct one. Rule~1 only matches a part of the 
    58       expression.\<close>
    59 
    60 end