src/sml/IsacKnowledge/Poly.thy
branchstart-work-070517
changeset 267 c02476bf9d9b
equal deleted inserted replaced
266:9acb256f8a40 267:c02476bf9d9b
       
     1 (* WN.020812: theorems in the Reals,
       
     2    necessary for special rule sets, in addition to Isabelle2002.
       
     3    !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
       
     4    !!! THIS IS THE _least_ NUMBER OF ADDITIONAL THEOREMS !!!
       
     5    !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
       
     6    xxxI contain ^^^ instead of ^ in the respective theorem xxx in 2002
       
     7    changed by: Richard Lang 020912
       
     8 *)
       
     9 
       
    10 (*
       
    11    use_thy"IsacKnowledge/Poly";
       
    12    use_thy"Poly";
       
    13    use_thy_only"IsacKnowledge/Poly";
       
    14 
       
    15    remove_thy"Poly";
       
    16    use_thy"IsacKnowledge/Isac";
       
    17 
       
    18 
       
    19    use"ROOT.ML";
       
    20    cd"IsacKnowledge";
       
    21  *)
       
    22 
       
    23 Poly = Simplify + 
       
    24 
       
    25 (*-------------------- consts-----------------------------------------------*)
       
    26 consts
       
    27 
       
    28   is'_expanded'_in :: "[real, real] => bool" ("_ is'_expanded'_in _") 
       
    29   is'_poly'_in :: "[real, real] => bool" ("_ is'_poly'_in _")          (*RL DA *)
       
    30   has'_degree'_in :: "[real, real] => real" ("_ has'_degree'_in _")(*RL DA *)
       
    31   is'_polyrat'_in :: "[real, real] => bool" ("_ is'_polyrat'_in _")(*RL030626*)
       
    32 
       
    33  is'_multUnordered  :: "real => bool" ("_ is'_multUnordered") 
       
    34  is'_addUnordered   :: "real => bool" ("_ is'_addUnordered") (*WN030618*)
       
    35  is'_polyexp        :: "real => bool" ("_ is'_polyexp") 
       
    36 
       
    37   Expand'_binoms
       
    38              :: "['y, \
       
    39 		  \ 'y] => 'y"
       
    40                ("((Script Expand'_binoms (_ =))// \
       
    41                  \ (_))" 9)
       
    42 
       
    43 (*-------------------- rules------------------------------------------------*)
       
    44 rules (*.not contained in Isabelle2002,
       
    45          stated as axioms, TODO: prove as theorems;
       
    46          theorem-IDs 'xxxI' with ^^^ instead of ^ in 'xxx' in Isabelle2002.*)
       
    47 
       
    48   realpow_pow             "(a ^^^ b) ^^^ c = a ^^^ (b * c)"
       
    49   realpow_addI            "r ^^^ (n + m) = r ^^^ n * r ^^^ m"
       
    50   realpow_addI_assoc_l    "r ^^^ n * (r ^^^ m * s) = r ^^^ (n + m) * s"
       
    51   realpow_addI_assoc_r    "s * r ^^^ n * r ^^^ m = s * r ^^^ (n + m)"
       
    52 		  
       
    53   realpow_oneI            "r ^^^ 1 = r"
       
    54   realpow_zeroI            "r ^^^ 0 = 1"
       
    55   realpow_eq_oneI         "1 ^^^ n = 1"
       
    56   realpow_multI           "(r * s) ^^^ n = r ^^^ n * s ^^^ n" 
       
    57   realpow_multI_poly      "[| r is_polyexp; s is_polyexp |] ==> \
       
    58 			      \(r * s) ^^^ n = r ^^^ n * s ^^^ n" 
       
    59   realpow_minus_oneI      "-1 ^^^ (2 * n) = 1"  
       
    60 
       
    61   realpow_twoI            "r ^^^ 2 = r * r"
       
    62   realpow_twoI_assoc_l	  "r * (r * s) = r ^^^ 2 * s"
       
    63   realpow_twoI_assoc_r	  "s * r * r = s * r ^^^ 2"
       
    64   realpow_two_atom        "r is_atom ==> r * r = r ^^^ 2"
       
    65   realpow_plus_1          "r * r ^^^ n = r ^^^ (n + 1)"         
       
    66   realpow_plus_1_assoc_l  "r * (r ^^^ m * s) = r ^^^ (1 + m) * s" 
       
    67   realpow_plus_1_assoc_l2 "r ^^^ m * (r * s) = r ^^^ (1 + m) * s" 
       
    68   realpow_plus_1_assoc_r  "s * r * r ^^^ m = s * r ^^^ (1 + m)"
       
    69   realpow_plus_1_atom     "r is_atom ==> r * r ^^^ n = r ^^^ (1 + n)"
       
    70   realpow_def_atom        "[| Not (r is_atom); 1 < n |] \
       
    71 			  \ ==> r ^^^ n = r * r ^^^ (n + -1)"
       
    72   realpow_addI_atom       "r is_atom ==> r ^^^ n * r ^^^ m = r ^^^ (n + m)"
       
    73 
       
    74 
       
    75   realpow_minus_even	  "n is_even ==> (- r) ^^^ n = r ^^^ n"
       
    76   realpow_minus_odd       "Not (n is_even) ==> (- r) ^^^ n = -1 * r ^^^ n"
       
    77 
       
    78 
       
    79 (* RL 020914 *)
       
    80   real_pp_binom_times        "(a + b)*(c + d) = a*c + a*d + b*c + b*d"
       
    81   real_pm_binom_times        "(a + b)*(c - d) = a*c - a*d + b*c - b*d"
       
    82   real_mp_binom_times        "(a - b)*(c + d) = a*c + a*d - b*c - b*d"
       
    83   real_mm_binom_times        "(a - b)*(c - d) = a*c - a*d - b*c + b*d"
       
    84   real_plus_binom_pow3       "(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3"
       
    85   real_plus_binom_pow3_poly  "[| a is_polyexp; b is_polyexp |] ==> \
       
    86 			      \(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3"
       
    87   real_minus_binom_pow3      "(a - b)^^^3 = a^^^3 - 3*a^^^2*b + 3*a*b^^^2 - b^^^3"
       
    88   real_minus_binom_pow3_p    "(a + -1 * b)^^^3 = a^^^3 + -3*a^^^2*b + 3*a*b^^^2 + -1*b^^^3"
       
    89 (* real_plus_binom_pow        "[| n is_const;  3 < n |] ==>  \
       
    90 			      \(a + b)^^^n = (a + b) * (a + b)^^^(n - 1)" *)
       
    91   real_plus_binom_pow4       "(a + b)^^^4 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)*(a + b)"
       
    92   real_plus_binom_pow4_poly  "[| a is_polyexp; b is_polyexp |] ==> \
       
    93 			      \(a + b)^^^4 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)*(a + b)"
       
    94   real_plus_binom_pow5       "(a + b)^^^5 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)*(a^^^2 + 2*a*b + b^^^2)"
       
    95 
       
    96   real_plus_binom_pow5_poly  "[| a is_polyexp; b is_polyexp |] ==> \
       
    97 			      \(a + b)^^^5 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)*(a^^^2 + 2*a*b + b^^^2)"
       
    98 
       
    99   real_diff_plus             "a - b = a + -b" (*17.3.03: do_NOT_use*)
       
   100   real_diff_minus            "a - b = a + -1 * b"
       
   101   real_plus_binom_times      "(a + b)*(a + b) = a^^^2 + 2*a*b + b^^^2"
       
   102   real_minus_binom_times     "(a - b)*(a - b) = a^^^2 - 2*a*b + b^^^2"
       
   103   (*WN071229 changed for Schaerding -----vvv*)
       
   104   (*real_plus_binom_pow2       "(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"*)
       
   105   real_plus_binom_pow2       "(a + b)^^^2 = (a + b) * (a + b)"
       
   106   (*WN071229 changed for Schaerding -----^^^*)
       
   107   real_plus_binom_pow2_poly   "[| a is_polyexp; b is_polyexp |] ==> \
       
   108 			      \(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"
       
   109   real_minus_binom_pow2      "(a - b)^^^2 = a^^^2 - 2*a*b + b^^^2"
       
   110   real_minus_binom_pow2_p    "(a - b)^^^2 = a^^^2 + -2*a*b + b^^^2"
       
   111   real_plus_minus_binom1     "(a + b)*(a - b) = a^^^2 - b^^^2"
       
   112   real_plus_minus_binom1_p   "(a + b)*(a - b) = a^^^2 + -1*b^^^2"
       
   113   real_plus_minus_binom1_p_p "(a + b)*(a + -1 * b) = a^^^2 + -1*b^^^2"
       
   114   real_plus_minus_binom2     "(a - b)*(a + b) = a^^^2 - b^^^2"
       
   115   real_plus_minus_binom2_p   "(a - b)*(a + b) = a^^^2 + -1*b^^^2"
       
   116   real_plus_minus_binom2_p_p "(a + -1 * b)*(a + b) = a^^^2 + -1*b^^^2"
       
   117   real_plus_binom_times1     "(a +  1*b)*(a + -1*b) = a^^^2 + -1*b^^^2"
       
   118   real_plus_binom_times2     "(a + -1*b)*(a +  1*b) = a^^^2 + -1*b^^^2"
       
   119 
       
   120   real_num_collect           "[| l is_const; m is_const |] ==> \
       
   121 					\l * n + m * n = (l + m) * n"
       
   122 (* FIXME.MG.0401: replace 'real_num_collect_assoc' 
       
   123 	by 'real_num_collect_assoc_l' ... are equal, introduced by MG ! *)
       
   124   real_num_collect_assoc     "[| l is_const; m is_const |] ==>  \
       
   125 					\l * n + (m * n + k) = (l + m) * n + k"
       
   126   real_num_collect_assoc_l     "[| l is_const; m is_const |] ==>  \
       
   127 					\l * n + (m * n + k) = (l + m)
       
   128 					* n + k"
       
   129   real_num_collect_assoc_r     "[| l is_const; m is_const |] ==>  \
       
   130 					\(k + m * n) + l * n = k + (l + m) * n"
       
   131   real_one_collect           "m is_const ==> n + m * n = (1 + m) * n"
       
   132 (* FIXME.MG.0401: replace 'real_one_collect_assoc' 
       
   133 	by 'real_one_collect_assoc_l' ... are equal, introduced by MG ! *)
       
   134   real_one_collect_assoc     "m is_const ==> n + (m * n + k) = (1 + m)* n + k"
       
   135 
       
   136   real_one_collect_assoc_l   "m is_const ==> n + (m * n + k) = (1 + m) * n + k"
       
   137   real_one_collect_assoc_r   "m is_const ==>(k + n) +  m * n = k + (1 + m) * n"
       
   138 
       
   139 (* FIXME.MG.0401: replace 'real_mult_2_assoc' 
       
   140 	by 'real_mult_2_assoc_l' ... are equal, introduced by MG ! *)
       
   141   real_mult_2_assoc          "z1 + (z1 + k) = 2 * z1 + k"
       
   142   real_mult_2_assoc_l        "z1 + (z1 + k) = 2 * z1 + k"
       
   143   real_mult_2_assoc_r        "(k + z1) + z1 = k + 2 * z1"
       
   144 
       
   145   real_add_mult_distrib_poly "w is_polyexp ==> (z1 + z2) * w = z1 * w + z2 * w"
       
   146   real_add_mult_distrib2_poly "w is_polyexp ==> w * (z1 + z2) = w * z1 + w * z2"
       
   147 end