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