src/HOL/Decision_Procs/Approximation.thy
changeset 30968 10fef94f40fc
parent 30952 7ab2716dd93b
child 30971 7fbebf75b3ef
equal deleted inserted replaced
30967:b5d67f83576e 30968:10fef94f40fc
    95   and ub_Suc: "\<And> n i k x. ub (Suc n) i k x = rapprox_rat prec 1 (int k) + x * (lb n (F i) (G i k) x)"
    95   and ub_Suc: "\<And> n i k x. ub (Suc n) i k x = rapprox_rat prec 1 (int k) + x * (lb n (F i) (G i k) x)"
    96   shows "Ifloat (lb n ((F o^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. (1 / real (f (j' + j))) * Ifloat x ^ j)" (is "?lb") and 
    96   shows "Ifloat (lb n ((F o^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. (1 / real (f (j' + j))) * Ifloat x ^ j)" (is "?lb") and 
    97     "(\<Sum>j=0..<n. (1 / real (f (j' + j))) * (Ifloat x ^ j)) \<le> Ifloat (ub n ((F o^ j') s) (f j') x)" (is "?ub")
    97     "(\<Sum>j=0..<n. (1 / real (f (j' + j))) * (Ifloat x ^ j)) \<le> Ifloat (ub n ((F o^ j') s) (f j') x)" (is "?ub")
    98 proof -
    98 proof -
    99   { fix x y z :: float have "x - y * z = x + - y * z"
    99   { fix x y z :: float have "x - y * z = x + - y * z"
   100       by (cases x, cases y, cases z, simp add: plus_float.simps minus_float.simps uminus_float.simps times_float.simps algebra_simps)
   100       by (cases x, cases y, cases z, simp add: plus_float.simps minus_float_def uminus_float.simps times_float.simps algebra_simps)
   101   } note diff_mult_minus = this
   101   } note diff_mult_minus = this
   102 
   102 
   103   { fix x :: float have "- (- x) = x" by (cases x, auto simp add: uminus_float.simps) } note minus_minus = this
   103   { fix x :: float have "- (- x) = x" by (cases x, auto simp add: uminus_float.simps) } note minus_minus = this
   104 
   104 
   105   have move_minus: "Ifloat (-x) = -1 * Ifloat x" by auto
   105   have move_minus: "Ifloat (-x) = -1 * Ifloat x" by auto
  1460     have "0 < Ifloat (?horner x) ^ num" using `0 < ?horner x`[unfolded less_float_def Ifloat_0] by (rule zero_less_power)
  1460     have "0 < Ifloat (?horner x) ^ num" using `0 < ?horner x`[unfolded less_float_def Ifloat_0] by (rule zero_less_power)
  1461     also have "\<dots> = Ifloat ((?horner x) ^ num)" using float_power by auto
  1461     also have "\<dots> = Ifloat ((?horner x) ^ num)" using float_power by auto
  1462     finally have "0 < Ifloat ((?horner x) ^ num)" .
  1462     finally have "0 < Ifloat ((?horner x) ^ num)" .
  1463   }
  1463   }
  1464   ultimately show ?thesis
  1464   ultimately show ?thesis
  1465     unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] Let_def by (cases "floor_fl x", cases "x < - 1", auto simp add: le_float_def less_float_def normfloat) 
  1465     unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] Let_def
       
  1466     by (cases "floor_fl x", cases "x < - 1", auto simp add: float_power le_float_def less_float_def)
  1466 qed
  1467 qed
  1467 
  1468 
  1468 lemma exp_boundaries': assumes "x \<le> 0"
  1469 lemma exp_boundaries': assumes "x \<le> 0"
  1469   shows "exp (Ifloat x) \<in> { Ifloat (lb_exp prec x) .. Ifloat (ub_exp prec x)}"
  1470   shows "exp (Ifloat x) \<in> { Ifloat (lb_exp prec x) .. Ifloat (ub_exp prec x)}"
  1470 proof -
  1471 proof -