doc-src/isac/jrocnik/Test_Complex.thy
author Walther Neuper <neuper@ist.tugraz.at>
Mon, 25 Jul 2011 17:44:19 +0200
branchdecompose-isar
changeset 42185 332a0653d4c9
parent 42164 dc2fe21d2183
child 42207 d18da7deb0bb
permissions -rw-r--r--
intermed. updated test + Test_Z_Transform.thy etc
     1 theory Test_Complex imports Isac begin
     2 
     3 section {*terms and operations*}
     4 ML {*
     5 print_depth 999; @{term "Complex 1 (2::real)"};
     6 *}
     7 ML {*
     8 @{term "Complex 1 2 + Complex 3 4"};
     9 print_depth 999; str2term "Complex 1 2 + Complex 3 4"; print_depth 999;
    10 *}
    11 ML {*
    12 @{term "Complex 1 2 * Complex 3 4"};
    13 @{term "Complex 1 2 - Complex 3 4"};
    14 @{term "Complex 1 2 / Complex 3 4"};
    15 (*@{term "Complex 1 2 ^ Complex 3 4"};
    16   Type unification failed: Clash of types "complex" and "nat"
    17   Operator:  op ^ (Complex 1 2) :: nat \<Rightarrow> complex
    18   Operand:   Complex 3 4 :: complex*)
    19 *}
    20 ML {*
    21 term2str @{term "Complex 1 2 + Complex 3 4"};
    22 term2str @{term "Complex 1 2 / Complex 3 4"};
    23 *}
    24 
    25 ML {*
    26 val a = @{term "Complex 1 2"};
    27 atomty a;
    28 val a = str2term "Complex 1 2";
    29 atomty a;
    30 *}
    31 ML {*
    32 val b = @{term "Complex 3 4"};
    33 val b = str2term "Complex 3 4";
    34 (*a + (b::complex); ERROR: term + term*)
    35 *}
    36 
    37 section {*use the operations for simplification*}
    38 subsection {*example 1*}
    39 ML {*
    40 print_depth 1;
    41 val (thy, ro, er) = (@{theory}, tless_true, eval_rls);
    42 *}
    43 ML {*
    44 val thm = @{thm "complex_add"};
    45 val t = str2term "Complex 1 2 + Complex 3 4";
    46 val SOME _ = rewrite_ thy ro er true thm t;
    47 val SOME (t', _) = rewrite_ thy ro er true thm t;
    48 "Complex (1 + 3) (2 + 4)" = term2str t';
    49 *}
    50 
    51 subsection {*example 2*}
    52 ML {*
    53 val Simplify_complex = append_rls "Simplify_complex" e_rls
    54   [ Thm  ("complex_add",num_str @{thm complex_add}),
    55 		    Thm  ("complex_mult",num_str @{thm complex_mult}),
    56 		    Rls_ norm_Poly
    57 		  ];
    58 *}
    59 ML {*
    60 val t = str2term "(Complex 1 2 * (Complex 3 4 + Complex 5 (6::real)))";
    61 term2str t = "Complex 1 2 * (Complex 3 4 + Complex 5 6)";
    62 atomty t;
    63 *}
    64 ML {*
    65 trace_rewrite := true;
    66 val SOME (t, _) = rewrite_set_ thy true Simplify_complex t;
    67 trace_rewrite := false;
    68 term2str t = "Complex -12 26";
    69 atomty t;
    70 *}
    71 ML {*
    72 
    73 *}
    74 ML {*
    75 trace_rewrite := true;
    76 trace_rewrite := false;
    77 *}
    78 
    79 end
    80