--- a/src/HOL/IsaMakefile Wed Sep 02 21:33:16 2009 +0200
+++ b/src/HOL/IsaMakefile Wed Sep 02 21:34:13 2009 +0200
@@ -941,7 +941,7 @@
HOL-Matrix: HOL $(LOG)/HOL-Matrix.gz
-$(LOG)/HOL-Matrix.gz: $(OUT)/HOL \
+$(LOG)/HOL-Matrix.gz: $(OUT)/HOL \
$(SRC)/Tools/Compute_Oracle/Compute_Oracle.thy \
$(SRC)/Tools/Compute_Oracle/am_compiler.ML \
$(SRC)/Tools/Compute_Oracle/am_interpreter.ML \
@@ -949,8 +949,8 @@
$(SRC)/Tools/Compute_Oracle/linker.ML \
$(SRC)/Tools/Compute_Oracle/am_ghc.ML \
$(SRC)/Tools/Compute_Oracle/am_sml.ML \
- $(SRC)/Tools/Compute_Oracle/compute.ML \
- Tools/ComputeFloat.thy Tools/float_arith.ML \
+ $(SRC)/Tools/Compute_Oracle/compute.ML Matrix/ComputeFloat.thy \
+ Matrix/ComputeHOL.thy Matrix/ComputeNumeral.thy Tools/float_arith.ML \
Matrix/Matrix.thy Matrix/SparseMatrix.thy Matrix/LP.thy \
Matrix/document/root.tex Matrix/ROOT.ML Matrix/cplex/Cplex.thy \
Matrix/cplex/CplexMatrixConverter.ML Matrix/cplex/Cplex_tools.ML \
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Matrix/ComputeFloat.thy Wed Sep 02 21:34:13 2009 +0200
@@ -0,0 +1,568 @@
+(* Title: HOL/Tools/ComputeFloat.thy
+ Author: Steven Obua
+*)
+
+header {* Floating Point Representation of the Reals *}
+
+theory ComputeFloat
+imports Complex_Main
+uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML")
+begin
+
+definition
+ pow2 :: "int \<Rightarrow> real" where
+ "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
+
+definition
+ float :: "int * int \<Rightarrow> real" where
+ "float x = real (fst x) * pow2 (snd x)"
+
+lemma pow2_0[simp]: "pow2 0 = 1"
+by (simp add: pow2_def)
+
+lemma pow2_1[simp]: "pow2 1 = 2"
+by (simp add: pow2_def)
+
+lemma pow2_neg: "pow2 x = inverse (pow2 (-x))"
+by (simp add: pow2_def)
+
+lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)"
+proof -
+ have h: "! n. nat (2 + int n) - Suc 0 = nat (1 + int n)" by arith
+ have g: "! a b. a - -1 = a + (1::int)" by arith
+ have pos: "! n. pow2 (int n + 1) = 2 * pow2 (int n)"
+ apply (auto, induct_tac n)
+ apply (simp_all add: pow2_def)
+ apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if])
+ by (auto simp add: h)
+ show ?thesis
+ proof (induct a)
+ case (1 n)
+ from pos show ?case by (simp add: algebra_simps)
+ next
+ case (2 n)
+ show ?case
+ apply (auto)
+ apply (subst pow2_neg[of "- int n"])
+ apply (subst pow2_neg[of "-1 - int n"])
+ apply (auto simp add: g pos)
+ done
+ qed
+qed
+
+lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)"
+proof (induct b)
+ case (1 n)
+ show ?case
+ proof (induct n)
+ case 0
+ show ?case by simp
+ next
+ case (Suc m)
+ show ?case by (auto simp add: algebra_simps pow2_add1 prems)
+ qed
+next
+ case (2 n)
+ show ?case
+ proof (induct n)
+ case 0
+ show ?case
+ apply (auto)
+ apply (subst pow2_neg[of "a + -1"])
+ apply (subst pow2_neg[of "-1"])
+ apply (simp)
+ apply (insert pow2_add1[of "-a"])
+ apply (simp add: algebra_simps)
+ apply (subst pow2_neg[of "-a"])
+ apply (simp)
+ done
+ case (Suc m)
+ have a: "int m - (a + -2) = 1 + (int m - a + 1)" by arith
+ have b: "int m - -2 = 1 + (int m + 1)" by arith
+ show ?case
+ apply (auto)
+ apply (subst pow2_neg[of "a + (-2 - int m)"])
+ apply (subst pow2_neg[of "-2 - int m"])
+ apply (auto simp add: algebra_simps)
+ apply (subst a)
+ apply (subst b)
+ apply (simp only: pow2_add1)
+ apply (subst pow2_neg[of "int m - a + 1"])
+ apply (subst pow2_neg[of "int m + 1"])
+ apply auto
+ apply (insert prems)
+ apply (auto simp add: algebra_simps)
+ done
+ qed
+qed
+
+lemma "float (a, e) + float (b, e) = float (a + b, e)"
+by (simp add: float_def algebra_simps)
+
+definition
+ int_of_real :: "real \<Rightarrow> int" where
+ "int_of_real x = (SOME y. real y = x)"
+
+definition
+ real_is_int :: "real \<Rightarrow> bool" where
+ "real_is_int x = (EX (u::int). x = real u)"
+
+lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))"
+by (auto simp add: real_is_int_def int_of_real_def)
+
+lemma float_transfer: "real_is_int ((real a)*(pow2 c)) \<Longrightarrow> float (a, b) = float (int_of_real ((real a)*(pow2 c)), b - c)"
+by (simp add: float_def real_is_int_def2 pow2_add[symmetric])
+
+lemma pow2_int: "pow2 (int c) = 2^c"
+by (simp add: pow2_def)
+
+lemma float_transfer_nat: "float (a, b) = float (a * 2^c, b - int c)"
+by (simp add: float_def pow2_int[symmetric] pow2_add[symmetric])
+
+lemma real_is_int_real[simp]: "real_is_int (real (x::int))"
+by (auto simp add: real_is_int_def int_of_real_def)
+
+lemma int_of_real_real[simp]: "int_of_real (real x) = x"
+by (simp add: int_of_real_def)
+
+lemma real_int_of_real[simp]: "real_is_int x \<Longrightarrow> real (int_of_real x) = x"
+by (auto simp add: int_of_real_def real_is_int_def)
+
+lemma real_is_int_add_int_of_real: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)"
+by (auto simp add: int_of_real_def real_is_int_def)
+
+lemma real_is_int_add[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a+b)"
+apply (subst real_is_int_def2)
+apply (simp add: real_is_int_add_int_of_real real_int_of_real)
+done
+
+lemma int_of_real_sub: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)"
+by (auto simp add: int_of_real_def real_is_int_def)
+
+lemma real_is_int_sub[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a-b)"
+apply (subst real_is_int_def2)
+apply (simp add: int_of_real_sub real_int_of_real)
+done
+
+lemma real_is_int_rep: "real_is_int x \<Longrightarrow> ?! (a::int). real a = x"
+by (auto simp add: real_is_int_def)
+
+lemma int_of_real_mult:
+ assumes "real_is_int a" "real_is_int b"
+ shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)"
+proof -
+ from prems have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto)
+ from prems have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto)
+ from a obtain a'::int where a':"a = real a'" by auto
+ from b obtain b'::int where b':"b = real b'" by auto
+ have r: "real a' * real b' = real (a' * b')" by auto
+ show ?thesis
+ apply (simp add: a' b')
+ apply (subst r)
+ apply (simp only: int_of_real_real)
+ done
+qed
+
+lemma real_is_int_mult[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a*b)"
+apply (subst real_is_int_def2)
+apply (simp add: int_of_real_mult)
+done
+
+lemma real_is_int_0[simp]: "real_is_int (0::real)"
+by (simp add: real_is_int_def int_of_real_def)
+
+lemma real_is_int_1[simp]: "real_is_int (1::real)"
+proof -
+ have "real_is_int (1::real) = real_is_int(real (1::int))" by auto
+ also have "\<dots> = True" by (simp only: real_is_int_real)
+ ultimately show ?thesis by auto
+qed
+
+lemma real_is_int_n1: "real_is_int (-1::real)"
+proof -
+ have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
+ also have "\<dots> = True" by (simp only: real_is_int_real)
+ ultimately show ?thesis by auto
+qed
+
+lemma real_is_int_number_of[simp]: "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
+proof -
+ have neg1: "real_is_int (-1::real)"
+ proof -
+ have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
+ also have "\<dots> = True" by (simp only: real_is_int_real)
+ ultimately show ?thesis by auto
+ qed
+
+ {
+ fix x :: int
+ have "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
+ unfolding number_of_eq
+ apply (induct x)
+ apply (induct_tac n)
+ apply (simp)
+ apply (simp)
+ apply (induct_tac n)
+ apply (simp add: neg1)
+ proof -
+ fix n :: nat
+ assume rn: "(real_is_int (of_int (- (int (Suc n)))))"
+ have s: "-(int (Suc (Suc n))) = -1 + - (int (Suc n))" by simp
+ show "real_is_int (of_int (- (int (Suc (Suc n)))))"
+ apply (simp only: s of_int_add)
+ apply (rule real_is_int_add)
+ apply (simp add: neg1)
+ apply (simp only: rn)
+ done
+ qed
+ }
+ note Abs_Bin = this
+ {
+ fix x :: int
+ have "? u. x = u"
+ apply (rule exI[where x = "x"])
+ apply (simp)
+ done
+ }
+ then obtain u::int where "x = u" by auto
+ with Abs_Bin show ?thesis by auto
+qed
+
+lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)"
+by (simp add: int_of_real_def)
+
+lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)"
+proof -
+ have 1: "(1::real) = real (1::int)" by auto
+ show ?thesis by (simp only: 1 int_of_real_real)
+qed
+
+lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b"
+proof -
+ have "real_is_int (number_of b)" by simp
+ then have uu: "?! u::int. number_of b = real u" by (auto simp add: real_is_int_rep)
+ then obtain u::int where u:"number_of b = real u" by auto
+ have "number_of b = real ((number_of b)::int)"
+ by (simp add: number_of_eq real_of_int_def)
+ have ub: "number_of b = real ((number_of b)::int)"
+ by (simp add: number_of_eq real_of_int_def)
+ from uu u ub have unb: "u = number_of b"
+ by blast
+ have "int_of_real (number_of b) = u" by (simp add: u)
+ with unb show ?thesis by simp
+qed
+
+lemma float_transfer_even: "even a \<Longrightarrow> float (a, b) = float (a div 2, b+1)"
+ apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified])
+ apply (simp_all add: pow2_def even_def real_is_int_def algebra_simps)
+ apply (auto)
+proof -
+ fix q::int
+ have a:"b - (-1\<Colon>int) = (1\<Colon>int) + b" by arith
+ show "(float (q, (b - (-1\<Colon>int)))) = (float (q, ((1\<Colon>int) + b)))"
+ by (simp add: a)
+qed
+
+lemma int_div_zdiv: "int (a div b) = (int a) div (int b)"
+by (rule zdiv_int)
+
+lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)"
+by (rule zmod_int)
+
+lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
+by arith
+
+function norm_float :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
+ "norm_float a b = (if a \<noteq> 0 \<and> even a then norm_float (a div 2) (b + 1)
+ else if a = 0 then (0, 0) else (a, b))"
+by auto
+
+termination by (relation "measure (nat o abs o fst)")
+ (auto intro: abs_div_2_less)
+
+lemma norm_float: "float x = float (split norm_float x)"
+proof -
+ {
+ fix a b :: int
+ have norm_float_pair: "float (a, b) = float (norm_float a b)"
+ proof (induct a b rule: norm_float.induct)
+ case (1 u v)
+ show ?case
+ proof cases
+ assume u: "u \<noteq> 0 \<and> even u"
+ with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto
+ with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even)
+ then show ?thesis
+ apply (subst norm_float.simps)
+ apply (simp add: ind)
+ done
+ next
+ assume "~(u \<noteq> 0 \<and> even u)"
+ then show ?thesis
+ by (simp add: prems float_def)
+ qed
+ qed
+ }
+ note helper = this
+ have "? a b. x = (a,b)" by auto
+ then obtain a b where "x = (a, b)" by blast
+ then show ?thesis by (simp add: helper)
+qed
+
+lemma float_add_l0: "float (0, e) + x = x"
+ by (simp add: float_def)
+
+lemma float_add_r0: "x + float (0, e) = x"
+ by (simp add: float_def)
+
+lemma float_add:
+ "float (a1, e1) + float (a2, e2) =
+ (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1)
+ else float (a1*2^(nat (e1-e2))+a2, e2))"
+ apply (simp add: float_def algebra_simps)
+ apply (auto simp add: pow2_int[symmetric] pow2_add[symmetric])
+ done
+
+lemma float_add_assoc1:
+ "(x + float (y1, e1)) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
+ by simp
+
+lemma float_add_assoc2:
+ "(float (y1, e1) + x) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
+ by simp
+
+lemma float_add_assoc3:
+ "float (y1, e1) + (x + float (y2, e2)) = (float (y1, e1) + float (y2, e2)) + x"
+ by simp
+
+lemma float_add_assoc4:
+ "float (y1, e1) + (float (y2, e2) + x) = (float (y1, e1) + float (y2, e2)) + x"
+ by simp
+
+lemma float_mult_l0: "float (0, e) * x = float (0, 0)"
+ by (simp add: float_def)
+
+lemma float_mult_r0: "x * float (0, e) = float (0, 0)"
+ by (simp add: float_def)
+
+definition
+ lbound :: "real \<Rightarrow> real"
+where
+ "lbound x = min 0 x"
+
+definition
+ ubound :: "real \<Rightarrow> real"
+where
+ "ubound x = max 0 x"
+
+lemma lbound: "lbound x \<le> x"
+ by (simp add: lbound_def)
+
+lemma ubound: "x \<le> ubound x"
+ by (simp add: ubound_def)
+
+lemma float_mult:
+ "float (a1, e1) * float (a2, e2) =
+ (float (a1 * a2, e1 + e2))"
+ by (simp add: float_def pow2_add)
+
+lemma float_minus:
+ "- (float (a,b)) = float (-a, b)"
+ by (simp add: float_def)
+
+lemma zero_less_pow2:
+ "0 < pow2 x"
+proof -
+ {
+ fix y
+ have "0 <= y \<Longrightarrow> 0 < pow2 y"
+ by (induct y, induct_tac n, simp_all add: pow2_add)
+ }
+ note helper=this
+ show ?thesis
+ apply (case_tac "0 <= x")
+ apply (simp add: helper)
+ apply (subst pow2_neg)
+ apply (simp add: helper)
+ done
+qed
+
+lemma zero_le_float:
+ "(0 <= float (a,b)) = (0 <= a)"
+ apply (auto simp add: float_def)
+ apply (auto simp add: zero_le_mult_iff zero_less_pow2)
+ apply (insert zero_less_pow2[of b])
+ apply (simp_all)
+ done
+
+lemma float_le_zero:
+ "(float (a,b) <= 0) = (a <= 0)"
+ apply (auto simp add: float_def)
+ apply (auto simp add: mult_le_0_iff)
+ apply (insert zero_less_pow2[of b])
+ apply auto
+ done
+
+lemma float_abs:
+ "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))"
+ apply (auto simp add: abs_if)
+ apply (simp_all add: zero_le_float[symmetric, of a b] float_minus)
+ done
+
+lemma float_zero:
+ "float (0, b) = 0"
+ by (simp add: float_def)
+
+lemma float_pprt:
+ "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))"
+ by (auto simp add: zero_le_float float_le_zero float_zero)
+
+lemma pprt_lbound: "pprt (lbound x) = float (0, 0)"
+ apply (simp add: float_def)
+ apply (rule pprt_eq_0)
+ apply (simp add: lbound_def)
+ done
+
+lemma nprt_ubound: "nprt (ubound x) = float (0, 0)"
+ apply (simp add: float_def)
+ apply (rule nprt_eq_0)
+ apply (simp add: ubound_def)
+ done
+
+lemma float_nprt:
+ "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))"
+ by (auto simp add: zero_le_float float_le_zero float_zero)
+
+lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1"
+ by auto
+
+lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)"
+ by simp
+
+lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)"
+ by simp
+
+lemma mult_left_one: "1 * a = (a::'a::semiring_1)"
+ by simp
+
+lemma mult_right_one: "a * 1 = (a::'a::semiring_1)"
+ by simp
+
+lemma int_pow_0: "(a::int)^(Numeral0) = 1"
+ by simp
+
+lemma int_pow_1: "(a::int)^(Numeral1) = a"
+ by simp
+
+lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0"
+ by simp
+
+lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1"
+ by simp
+
+lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0"
+ by simp
+
+lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1"
+ by simp
+
+lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1"
+ by simp
+
+lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1"
+proof -
+ have 1:"((-1)::nat) = 0"
+ by simp
+ show ?thesis by (simp add: 1)
+qed
+
+lemma fst_cong: "a=a' \<Longrightarrow> fst (a,b) = fst (a',b)"
+ by simp
+
+lemma snd_cong: "b=b' \<Longrightarrow> snd (a,b) = snd (a,b')"
+ by simp
+
+lemma lift_bool: "x \<Longrightarrow> x=True"
+ by simp
+
+lemma nlift_bool: "~x \<Longrightarrow> x=False"
+ by simp
+
+lemma not_false_eq_true: "(~ False) = True" by simp
+
+lemma not_true_eq_false: "(~ True) = False" by simp
+
+lemmas binarith =
+ normalize_bin_simps
+ pred_bin_simps succ_bin_simps
+ add_bin_simps minus_bin_simps mult_bin_simps
+
+lemma int_eq_number_of_eq:
+ "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)"
+ by (rule eq_number_of_eq)
+
+lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)"
+ by (simp only: iszero_number_of_Pls)
+
+lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))"
+ by simp
+
+lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)"
+ by simp
+
+lemma int_iszero_number_of_Bit1: "\<not> iszero ((number_of (Int.Bit1 w))::int)"
+ by simp
+
+lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)"
+ unfolding neg_def number_of_is_id by simp
+
+lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))"
+ by simp
+
+lemma int_neg_number_of_Min: "neg (-1::int)"
+ by simp
+
+lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)"
+ by simp
+
+lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)"
+ by simp
+
+lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (y + (uminus x)))::int))"
+ unfolding neg_def number_of_is_id by (simp add: not_less)
+
+lemmas intarithrel =
+ int_eq_number_of_eq
+ lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0
+ lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min]
+ int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq
+
+lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)"
+ by simp
+
+lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))"
+ by simp
+
+lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)"
+ by simp
+
+lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)"
+ by simp
+
+lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym
+
+lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of
+
+lemmas powerarith = nat_number_of zpower_number_of_even
+ zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
+ zpower_Pls zpower_Min
+
+lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0
+ float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound
+
+(* for use with the compute oracle *)
+lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false
+
+use "~~/src/HOL/Tools/float_arith.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Matrix/ComputeHOL.thy Wed Sep 02 21:34:13 2009 +0200
@@ -0,0 +1,191 @@
+theory ComputeHOL
+imports Complex_Main "~~/src/Tools/Compute_Oracle/Compute_Oracle"
+begin
+
+lemma Trueprop_eq_eq: "Trueprop X == (X == True)" by (simp add: atomize_eq)
+lemma meta_eq_trivial: "x == y \<Longrightarrow> x == y" by simp
+lemma meta_eq_imp_eq: "x == y \<Longrightarrow> x = y" by auto
+lemma eq_trivial: "x = y \<Longrightarrow> x = y" by auto
+lemma bool_to_true: "x :: bool \<Longrightarrow> x == True" by simp
+lemma transmeta_1: "x = y \<Longrightarrow> y == z \<Longrightarrow> x = z" by simp
+lemma transmeta_2: "x == y \<Longrightarrow> y = z \<Longrightarrow> x = z" by simp
+lemma transmeta_3: "x == y \<Longrightarrow> y == z \<Longrightarrow> x = z" by simp
+
+
+(**** compute_if ****)
+
+lemma If_True: "If True = (\<lambda> x y. x)" by ((rule ext)+,auto)
+lemma If_False: "If False = (\<lambda> x y. y)" by ((rule ext)+, auto)
+
+lemmas compute_if = If_True If_False
+
+(**** compute_bool ****)
+
+lemma bool1: "(\<not> True) = False" by blast
+lemma bool2: "(\<not> False) = True" by blast
+lemma bool3: "(P \<and> True) = P" by blast
+lemma bool4: "(True \<and> P) = P" by blast
+lemma bool5: "(P \<and> False) = False" by blast
+lemma bool6: "(False \<and> P) = False" by blast
+lemma bool7: "(P \<or> True) = True" by blast
+lemma bool8: "(True \<or> P) = True" by blast
+lemma bool9: "(P \<or> False) = P" by blast
+lemma bool10: "(False \<or> P) = P" by blast
+lemma bool11: "(True \<longrightarrow> P) = P" by blast
+lemma bool12: "(P \<longrightarrow> True) = True" by blast
+lemma bool13: "(True \<longrightarrow> P) = P" by blast
+lemma bool14: "(P \<longrightarrow> False) = (\<not> P)" by blast
+lemma bool15: "(False \<longrightarrow> P) = True" by blast
+lemma bool16: "(False = False) = True" by blast
+lemma bool17: "(True = True) = True" by blast
+lemma bool18: "(False = True) = False" by blast
+lemma bool19: "(True = False) = False" by blast
+
+lemmas compute_bool = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10 bool11 bool12 bool13 bool14 bool15 bool16 bool17 bool18 bool19
+
+
+(*** compute_pair ***)
+
+lemma compute_fst: "fst (x,y) = x" by simp
+lemma compute_snd: "snd (x,y) = y" by simp
+lemma compute_pair_eq: "((a, b) = (c, d)) = (a = c \<and> b = d)" by auto
+
+lemma prod_case_simp: "prod_case f (x,y) = f x y" by simp
+
+lemmas compute_pair = compute_fst compute_snd compute_pair_eq prod_case_simp
+
+(*** compute_option ***)
+
+lemma compute_the: "the (Some x) = x" by simp
+lemma compute_None_Some_eq: "(None = Some x) = False" by auto
+lemma compute_Some_None_eq: "(Some x = None) = False" by auto
+lemma compute_None_None_eq: "(None = None) = True" by auto
+lemma compute_Some_Some_eq: "(Some x = Some y) = (x = y)" by auto
+
+definition
+ option_case_compute :: "'b option \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+ "option_case_compute opt a f = option_case a f opt"
+
+lemma option_case_compute: "option_case = (\<lambda> a f opt. option_case_compute opt a f)"
+ by (simp add: option_case_compute_def)
+
+lemma option_case_compute_None: "option_case_compute None = (\<lambda> a f. a)"
+ apply (rule ext)+
+ apply (simp add: option_case_compute_def)
+ done
+
+lemma option_case_compute_Some: "option_case_compute (Some x) = (\<lambda> a f. f x)"
+ apply (rule ext)+
+ apply (simp add: option_case_compute_def)
+ done
+
+lemmas compute_option_case = option_case_compute option_case_compute_None option_case_compute_Some
+
+lemmas compute_option = compute_the compute_None_Some_eq compute_Some_None_eq compute_None_None_eq compute_Some_Some_eq compute_option_case
+
+(**** compute_list_length ****)
+
+lemma length_cons:"length (x#xs) = 1 + (length xs)"
+ by simp
+
+lemma length_nil: "length [] = 0"
+ by simp
+
+lemmas compute_list_length = length_nil length_cons
+
+(*** compute_list_case ***)
+
+definition
+ list_case_compute :: "'b list \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b list \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+ "list_case_compute l a f = list_case a f l"
+
+lemma list_case_compute: "list_case = (\<lambda> (a::'a) f (l::'b list). list_case_compute l a f)"
+ apply (rule ext)+
+ apply (simp add: list_case_compute_def)
+ done
+
+lemma list_case_compute_empty: "list_case_compute ([]::'b list) = (\<lambda> (a::'a) f. a)"
+ apply (rule ext)+
+ apply (simp add: list_case_compute_def)
+ done
+
+lemma list_case_compute_cons: "list_case_compute (u#v) = (\<lambda> (a::'a) f. (f (u::'b) v))"
+ apply (rule ext)+
+ apply (simp add: list_case_compute_def)
+ done
+
+lemmas compute_list_case = list_case_compute list_case_compute_empty list_case_compute_cons
+
+(*** compute_list_nth ***)
+(* Of course, you will need computation with nats for this to work \<dots> *)
+
+lemma compute_list_nth: "((x#xs) ! n) = (if n = 0 then x else (xs ! (n - 1)))"
+ by (cases n, auto)
+
+(*** compute_list ***)
+
+lemmas compute_list = compute_list_case compute_list_length compute_list_nth
+
+(*** compute_let ***)
+
+lemmas compute_let = Let_def
+
+(***********************)
+(* Everything together *)
+(***********************)
+
+lemmas compute_hol = compute_if compute_bool compute_pair compute_option compute_list compute_let
+
+ML {*
+signature ComputeHOL =
+sig
+ val prep_thms : thm list -> thm list
+ val to_meta_eq : thm -> thm
+ val to_hol_eq : thm -> thm
+ val symmetric : thm -> thm
+ val trans : thm -> thm -> thm
+end
+
+structure ComputeHOL : ComputeHOL =
+struct
+
+local
+fun lhs_of eq = fst (Thm.dest_equals (cprop_of eq));
+in
+fun rewrite_conv [] ct = raise CTERM ("rewrite_conv", [])
+ | rewrite_conv (eq :: eqs) ct =
+ Thm.instantiate (Thm.match (lhs_of eq, ct)) eq
+ handle Pattern.MATCH => rewrite_conv eqs ct;
+end
+
+val convert_conditions = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.try_conv (rewrite_conv [@{thm "Trueprop_eq_eq"}])))
+
+val eq_th = @{thm "HOL.eq_reflection"}
+val meta_eq_trivial = @{thm "ComputeHOL.meta_eq_trivial"}
+val bool_to_true = @{thm "ComputeHOL.bool_to_true"}
+
+fun to_meta_eq th = eq_th OF [th] handle THM _ => meta_eq_trivial OF [th] handle THM _ => bool_to_true OF [th]
+
+fun to_hol_eq th = @{thm "meta_eq_imp_eq"} OF [th] handle THM _ => @{thm "eq_trivial"} OF [th]
+
+fun prep_thms ths = map (convert_conditions o to_meta_eq) ths
+
+fun symmetric th = @{thm "HOL.sym"} OF [th] handle THM _ => @{thm "Pure.symmetric"} OF [th]
+
+local
+ val trans_HOL = @{thm "HOL.trans"}
+ val trans_HOL_1 = @{thm "ComputeHOL.transmeta_1"}
+ val trans_HOL_2 = @{thm "ComputeHOL.transmeta_2"}
+ val trans_HOL_3 = @{thm "ComputeHOL.transmeta_3"}
+ fun tr [] th1 th2 = trans_HOL OF [th1, th2]
+ | tr (t::ts) th1 th2 = (t OF [th1, th2] handle THM _ => tr ts th1 th2)
+in
+ fun trans th1 th2 = tr [trans_HOL, trans_HOL_1, trans_HOL_2, trans_HOL_3] th1 th2
+end
+
+end
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Matrix/ComputeNumeral.thy Wed Sep 02 21:34:13 2009 +0200
@@ -0,0 +1,195 @@
+theory ComputeNumeral
+imports ComputeHOL ComputeFloat
+begin
+
+(* normalization of bit strings *)
+lemmas bitnorm = normalize_bin_simps
+
+(* neg for bit strings *)
+lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def)
+lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto
+lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto
+lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto
+lemmas bitneg = neg1 neg2 neg3 neg4
+
+(* iszero for bit strings *)
+lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def)
+lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp
+lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto
+lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+ apply simp by arith
+lemmas bitiszero = iszero1 iszero2 iszero3 iszero4
+
+(* lezero for bit strings *)
+constdefs
+ "lezero x == (x \<le> 0)"
+lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto
+lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto
+lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto
+lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto
+lemmas bitlezero = lezero1 lezero2 lezero3 lezero4
+
+(* equality for bit strings *)
+lemmas biteq = eq_bin_simps
+
+(* x < y for bit strings *)
+lemmas bitless = less_bin_simps
+
+(* x \<le> y for bit strings *)
+lemmas bitle = le_bin_simps
+
+(* succ for bit strings *)
+lemmas bitsucc = succ_bin_simps
+
+(* pred for bit strings *)
+lemmas bitpred = pred_bin_simps
+
+(* unary minus for bit strings *)
+lemmas bituminus = minus_bin_simps
+
+(* addition for bit strings *)
+lemmas bitadd = add_bin_simps
+
+(* multiplication for bit strings *)
+lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def)
+lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute, simp add: mult_Min)
+lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0)
+lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp
+lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)"
+ unfolding Bit0_def Bit1_def by (simp add: algebra_simps)
+lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1
+
+lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul
+
+constdefs
+ "nat_norm_number_of (x::nat) == x"
+
+lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)"
+ apply (simp add: nat_norm_number_of_def)
+ unfolding lezero_def iszero_def neg_def
+ apply (simp add: numeral_simps)
+ done
+
+(* Normalization of nat literals *)
+lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto
+lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)" by auto
+lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of
+
+(* Suc *)
+lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id)
+
+(* Addition for nat *)
+lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))"
+ unfolding nat_number_of_def number_of_is_id neg_def
+ by auto
+
+(* Subtraction for nat *)
+lemma natsub: "(number_of x) - ((number_of y)::nat) =
+ (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))"
+ unfolding nat_norm_number_of
+ by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def)
+
+(* Multiplication for nat *)
+lemma natmul: "(number_of x) * ((number_of y)::nat) =
+ (if neg x then 0 else (if neg y then 0 else number_of (x * y)))"
+ unfolding nat_number_of_def number_of_is_id neg_def
+ by (simp add: nat_mult_distrib)
+
+lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \<and> lezero y) \<or> (x = y))"
+ by (auto simp add: iszero_def lezero_def neg_def number_of_is_id)
+
+lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \<and> (\<not> (lezero y)))"
+ by (simp add: lezero_def numeral_simps not_le)
+
+lemma natle: "(((number_of x)::nat) \<le> (number_of y)) = (y < x \<longrightarrow> lezero x)"
+ by (auto simp add: number_of_is_id lezero_def nat_number_of_def)
+
+fun natfac :: "nat \<Rightarrow> nat"
+where
+ "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))"
+
+lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps
+
+lemma number_eq: "(((number_of x)::'a::{number_ring, ordered_idom}) = (number_of y)) = (x = y)"
+ unfolding number_of_eq
+ apply simp
+ done
+
+lemma number_le: "(((number_of x)::'a::{number_ring, ordered_idom}) \<le> (number_of y)) = (x \<le> y)"
+ unfolding number_of_eq
+ apply simp
+ done
+
+lemma number_less: "(((number_of x)::'a::{number_ring, ordered_idom}) < (number_of y)) = (x < y)"
+ unfolding number_of_eq
+ apply simp
+ done
+
+lemma number_diff: "((number_of x)::'a::{number_ring, ordered_idom}) - number_of y = number_of (x + (- y))"
+ apply (subst diff_number_of_eq)
+ apply simp
+ done
+
+lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric]
+
+lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less
+
+lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)"
+ by (simp only: real_of_nat_number_of number_of_is_id)
+
+lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)"
+ by simp
+
+lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of
+
+lemmas zpowerarith = zpower_number_of_even
+ zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
+ zpower_Pls zpower_Min
+
+(* div, mod *)
+
+lemma adjust: "adjust b (q, r) = (if 0 \<le> r - b then (2 * q + 1, r - b) else (2 * q, r))"
+ by (auto simp only: adjust_def)
+
+lemma negateSnd: "negateSnd (q, r) = (q, -r)"
+ by (simp add: negateSnd_def)
+
+lemma divmod: "IntDiv.divmod a b = (if 0\<le>a then
+ if 0\<le>b then posDivAlg a b
+ else if a=0 then (0, 0)
+ else negateSnd (negDivAlg (-a) (-b))
+ else
+ if 0<b then negDivAlg a b
+ else negateSnd (posDivAlg (-a) (-b)))"
+ by (auto simp only: IntDiv.divmod_def)
+
+lemmas compute_div_mod = div_def mod_def divmod adjust negateSnd posDivAlg.simps negDivAlg.simps
+
+
+
+(* collecting all the theorems *)
+
+lemma even_Pls: "even (Int.Pls) = True"
+ apply (unfold Pls_def even_def)
+ by simp
+
+lemma even_Min: "even (Int.Min) = False"
+ apply (unfold Min_def even_def)
+ by simp
+
+lemma even_B0: "even (Int.Bit0 x) = True"
+ apply (unfold Bit0_def)
+ by simp
+
+lemma even_B1: "even (Int.Bit1 x) = False"
+ apply (unfold Bit1_def)
+ by simp
+
+lemma even_number_of: "even ((number_of w)::int) = even w"
+ by (simp only: number_of_is_id)
+
+lemmas compute_even = even_Pls even_Min even_B0 even_B1 even_number_of
+
+lemmas compute_numeral = compute_if compute_let compute_pair compute_bool
+ compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even
+
+end
--- a/src/HOL/Matrix/LP.thy Wed Sep 02 21:33:16 2009 +0200
+++ b/src/HOL/Matrix/LP.thy Wed Sep 02 21:34:13 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Matrix/LP.thy
- ID: $Id$
Author: Steven Obua
*)
--- a/src/HOL/Matrix/Matrix.thy Wed Sep 02 21:33:16 2009 +0200
+++ b/src/HOL/Matrix/Matrix.thy Wed Sep 02 21:34:13 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Matrix/Matrix.thy
- ID: $Id$
Author: Steven Obua
*)
--- a/src/HOL/Matrix/ROOT.ML Wed Sep 02 21:33:16 2009 +0200
+++ b/src/HOL/Matrix/ROOT.ML Wed Sep 02 21:34:13 2009 +0200
@@ -1,5 +1,1 @@
-(* Title: HOL/Matrix/ROOT.ML
- ID: $Id$
-*)
-
use_thys ["SparseMatrix", "cplex/Cplex"];
--- a/src/HOL/Matrix/SparseMatrix.thy Wed Sep 02 21:33:16 2009 +0200
+++ b/src/HOL/Matrix/SparseMatrix.thy Wed Sep 02 21:34:13 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Matrix/SparseMatrix.thy
- ID: $Id$
Author: Steven Obua
*)
--- a/src/HOL/Matrix/cplex/Cplex.thy Wed Sep 02 21:33:16 2009 +0200
+++ b/src/HOL/Matrix/cplex/Cplex.thy Wed Sep 02 21:34:13 2009 +0200
@@ -1,10 +1,9 @@
(* Title: HOL/Matrix/cplex/Cplex.thy
- ID: $Id$
Author: Steven Obua
*)
theory Cplex
-imports SparseMatrix LP "~~/src/HOL/Tools/ComputeFloat" "~~/src/HOL/Tools/ComputeNumeral"
+imports SparseMatrix LP ComputeFloat ComputeNumeral
uses "Cplex_tools.ML" "CplexMatrixConverter.ML" "FloatSparseMatrixBuilder.ML"
"fspmlp.ML" ("matrixlp.ML")
begin
--- a/src/HOL/Matrix/cplex/CplexMatrixConverter.ML Wed Sep 02 21:33:16 2009 +0200
+++ b/src/HOL/Matrix/cplex/CplexMatrixConverter.ML Wed Sep 02 21:34:13 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Matrix/cplex/CplexMatrixConverter.ML
- ID: $Id$
Author: Steven Obua
*)
--- a/src/HOL/Matrix/cplex/Cplex_tools.ML Wed Sep 02 21:33:16 2009 +0200
+++ b/src/HOL/Matrix/cplex/Cplex_tools.ML Wed Sep 02 21:34:13 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Matrix/cplex/Cplex_tools.ML
- ID: $Id$
Author: Steven Obua
*)
--- a/src/HOL/Tools/ComputeFloat.thy Wed Sep 02 21:33:16 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,568 +0,0 @@
-(* Title: HOL/Tools/ComputeFloat.thy
- Author: Steven Obua
-*)
-
-header {* Floating Point Representation of the Reals *}
-
-theory ComputeFloat
-imports Complex_Main
-uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML")
-begin
-
-definition
- pow2 :: "int \<Rightarrow> real" where
- "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
-
-definition
- float :: "int * int \<Rightarrow> real" where
- "float x = real (fst x) * pow2 (snd x)"
-
-lemma pow2_0[simp]: "pow2 0 = 1"
-by (simp add: pow2_def)
-
-lemma pow2_1[simp]: "pow2 1 = 2"
-by (simp add: pow2_def)
-
-lemma pow2_neg: "pow2 x = inverse (pow2 (-x))"
-by (simp add: pow2_def)
-
-lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)"
-proof -
- have h: "! n. nat (2 + int n) - Suc 0 = nat (1 + int n)" by arith
- have g: "! a b. a - -1 = a + (1::int)" by arith
- have pos: "! n. pow2 (int n + 1) = 2 * pow2 (int n)"
- apply (auto, induct_tac n)
- apply (simp_all add: pow2_def)
- apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if])
- by (auto simp add: h)
- show ?thesis
- proof (induct a)
- case (1 n)
- from pos show ?case by (simp add: algebra_simps)
- next
- case (2 n)
- show ?case
- apply (auto)
- apply (subst pow2_neg[of "- int n"])
- apply (subst pow2_neg[of "-1 - int n"])
- apply (auto simp add: g pos)
- done
- qed
-qed
-
-lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)"
-proof (induct b)
- case (1 n)
- show ?case
- proof (induct n)
- case 0
- show ?case by simp
- next
- case (Suc m)
- show ?case by (auto simp add: algebra_simps pow2_add1 prems)
- qed
-next
- case (2 n)
- show ?case
- proof (induct n)
- case 0
- show ?case
- apply (auto)
- apply (subst pow2_neg[of "a + -1"])
- apply (subst pow2_neg[of "-1"])
- apply (simp)
- apply (insert pow2_add1[of "-a"])
- apply (simp add: algebra_simps)
- apply (subst pow2_neg[of "-a"])
- apply (simp)
- done
- case (Suc m)
- have a: "int m - (a + -2) = 1 + (int m - a + 1)" by arith
- have b: "int m - -2 = 1 + (int m + 1)" by arith
- show ?case
- apply (auto)
- apply (subst pow2_neg[of "a + (-2 - int m)"])
- apply (subst pow2_neg[of "-2 - int m"])
- apply (auto simp add: algebra_simps)
- apply (subst a)
- apply (subst b)
- apply (simp only: pow2_add1)
- apply (subst pow2_neg[of "int m - a + 1"])
- apply (subst pow2_neg[of "int m + 1"])
- apply auto
- apply (insert prems)
- apply (auto simp add: algebra_simps)
- done
- qed
-qed
-
-lemma "float (a, e) + float (b, e) = float (a + b, e)"
-by (simp add: float_def algebra_simps)
-
-definition
- int_of_real :: "real \<Rightarrow> int" where
- "int_of_real x = (SOME y. real y = x)"
-
-definition
- real_is_int :: "real \<Rightarrow> bool" where
- "real_is_int x = (EX (u::int). x = real u)"
-
-lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))"
-by (auto simp add: real_is_int_def int_of_real_def)
-
-lemma float_transfer: "real_is_int ((real a)*(pow2 c)) \<Longrightarrow> float (a, b) = float (int_of_real ((real a)*(pow2 c)), b - c)"
-by (simp add: float_def real_is_int_def2 pow2_add[symmetric])
-
-lemma pow2_int: "pow2 (int c) = 2^c"
-by (simp add: pow2_def)
-
-lemma float_transfer_nat: "float (a, b) = float (a * 2^c, b - int c)"
-by (simp add: float_def pow2_int[symmetric] pow2_add[symmetric])
-
-lemma real_is_int_real[simp]: "real_is_int (real (x::int))"
-by (auto simp add: real_is_int_def int_of_real_def)
-
-lemma int_of_real_real[simp]: "int_of_real (real x) = x"
-by (simp add: int_of_real_def)
-
-lemma real_int_of_real[simp]: "real_is_int x \<Longrightarrow> real (int_of_real x) = x"
-by (auto simp add: int_of_real_def real_is_int_def)
-
-lemma real_is_int_add_int_of_real: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)"
-by (auto simp add: int_of_real_def real_is_int_def)
-
-lemma real_is_int_add[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a+b)"
-apply (subst real_is_int_def2)
-apply (simp add: real_is_int_add_int_of_real real_int_of_real)
-done
-
-lemma int_of_real_sub: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)"
-by (auto simp add: int_of_real_def real_is_int_def)
-
-lemma real_is_int_sub[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a-b)"
-apply (subst real_is_int_def2)
-apply (simp add: int_of_real_sub real_int_of_real)
-done
-
-lemma real_is_int_rep: "real_is_int x \<Longrightarrow> ?! (a::int). real a = x"
-by (auto simp add: real_is_int_def)
-
-lemma int_of_real_mult:
- assumes "real_is_int a" "real_is_int b"
- shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)"
-proof -
- from prems have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto)
- from prems have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto)
- from a obtain a'::int where a':"a = real a'" by auto
- from b obtain b'::int where b':"b = real b'" by auto
- have r: "real a' * real b' = real (a' * b')" by auto
- show ?thesis
- apply (simp add: a' b')
- apply (subst r)
- apply (simp only: int_of_real_real)
- done
-qed
-
-lemma real_is_int_mult[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a*b)"
-apply (subst real_is_int_def2)
-apply (simp add: int_of_real_mult)
-done
-
-lemma real_is_int_0[simp]: "real_is_int (0::real)"
-by (simp add: real_is_int_def int_of_real_def)
-
-lemma real_is_int_1[simp]: "real_is_int (1::real)"
-proof -
- have "real_is_int (1::real) = real_is_int(real (1::int))" by auto
- also have "\<dots> = True" by (simp only: real_is_int_real)
- ultimately show ?thesis by auto
-qed
-
-lemma real_is_int_n1: "real_is_int (-1::real)"
-proof -
- have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
- also have "\<dots> = True" by (simp only: real_is_int_real)
- ultimately show ?thesis by auto
-qed
-
-lemma real_is_int_number_of[simp]: "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
-proof -
- have neg1: "real_is_int (-1::real)"
- proof -
- have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
- also have "\<dots> = True" by (simp only: real_is_int_real)
- ultimately show ?thesis by auto
- qed
-
- {
- fix x :: int
- have "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
- unfolding number_of_eq
- apply (induct x)
- apply (induct_tac n)
- apply (simp)
- apply (simp)
- apply (induct_tac n)
- apply (simp add: neg1)
- proof -
- fix n :: nat
- assume rn: "(real_is_int (of_int (- (int (Suc n)))))"
- have s: "-(int (Suc (Suc n))) = -1 + - (int (Suc n))" by simp
- show "real_is_int (of_int (- (int (Suc (Suc n)))))"
- apply (simp only: s of_int_add)
- apply (rule real_is_int_add)
- apply (simp add: neg1)
- apply (simp only: rn)
- done
- qed
- }
- note Abs_Bin = this
- {
- fix x :: int
- have "? u. x = u"
- apply (rule exI[where x = "x"])
- apply (simp)
- done
- }
- then obtain u::int where "x = u" by auto
- with Abs_Bin show ?thesis by auto
-qed
-
-lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)"
-by (simp add: int_of_real_def)
-
-lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)"
-proof -
- have 1: "(1::real) = real (1::int)" by auto
- show ?thesis by (simp only: 1 int_of_real_real)
-qed
-
-lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b"
-proof -
- have "real_is_int (number_of b)" by simp
- then have uu: "?! u::int. number_of b = real u" by (auto simp add: real_is_int_rep)
- then obtain u::int where u:"number_of b = real u" by auto
- have "number_of b = real ((number_of b)::int)"
- by (simp add: number_of_eq real_of_int_def)
- have ub: "number_of b = real ((number_of b)::int)"
- by (simp add: number_of_eq real_of_int_def)
- from uu u ub have unb: "u = number_of b"
- by blast
- have "int_of_real (number_of b) = u" by (simp add: u)
- with unb show ?thesis by simp
-qed
-
-lemma float_transfer_even: "even a \<Longrightarrow> float (a, b) = float (a div 2, b+1)"
- apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified])
- apply (simp_all add: pow2_def even_def real_is_int_def algebra_simps)
- apply (auto)
-proof -
- fix q::int
- have a:"b - (-1\<Colon>int) = (1\<Colon>int) + b" by arith
- show "(float (q, (b - (-1\<Colon>int)))) = (float (q, ((1\<Colon>int) + b)))"
- by (simp add: a)
-qed
-
-lemma int_div_zdiv: "int (a div b) = (int a) div (int b)"
-by (rule zdiv_int)
-
-lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)"
-by (rule zmod_int)
-
-lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
-by arith
-
-function norm_float :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
- "norm_float a b = (if a \<noteq> 0 \<and> even a then norm_float (a div 2) (b + 1)
- else if a = 0 then (0, 0) else (a, b))"
-by auto
-
-termination by (relation "measure (nat o abs o fst)")
- (auto intro: abs_div_2_less)
-
-lemma norm_float: "float x = float (split norm_float x)"
-proof -
- {
- fix a b :: int
- have norm_float_pair: "float (a, b) = float (norm_float a b)"
- proof (induct a b rule: norm_float.induct)
- case (1 u v)
- show ?case
- proof cases
- assume u: "u \<noteq> 0 \<and> even u"
- with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto
- with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even)
- then show ?thesis
- apply (subst norm_float.simps)
- apply (simp add: ind)
- done
- next
- assume "~(u \<noteq> 0 \<and> even u)"
- then show ?thesis
- by (simp add: prems float_def)
- qed
- qed
- }
- note helper = this
- have "? a b. x = (a,b)" by auto
- then obtain a b where "x = (a, b)" by blast
- then show ?thesis by (simp add: helper)
-qed
-
-lemma float_add_l0: "float (0, e) + x = x"
- by (simp add: float_def)
-
-lemma float_add_r0: "x + float (0, e) = x"
- by (simp add: float_def)
-
-lemma float_add:
- "float (a1, e1) + float (a2, e2) =
- (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1)
- else float (a1*2^(nat (e1-e2))+a2, e2))"
- apply (simp add: float_def algebra_simps)
- apply (auto simp add: pow2_int[symmetric] pow2_add[symmetric])
- done
-
-lemma float_add_assoc1:
- "(x + float (y1, e1)) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
- by simp
-
-lemma float_add_assoc2:
- "(float (y1, e1) + x) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
- by simp
-
-lemma float_add_assoc3:
- "float (y1, e1) + (x + float (y2, e2)) = (float (y1, e1) + float (y2, e2)) + x"
- by simp
-
-lemma float_add_assoc4:
- "float (y1, e1) + (float (y2, e2) + x) = (float (y1, e1) + float (y2, e2)) + x"
- by simp
-
-lemma float_mult_l0: "float (0, e) * x = float (0, 0)"
- by (simp add: float_def)
-
-lemma float_mult_r0: "x * float (0, e) = float (0, 0)"
- by (simp add: float_def)
-
-definition
- lbound :: "real \<Rightarrow> real"
-where
- "lbound x = min 0 x"
-
-definition
- ubound :: "real \<Rightarrow> real"
-where
- "ubound x = max 0 x"
-
-lemma lbound: "lbound x \<le> x"
- by (simp add: lbound_def)
-
-lemma ubound: "x \<le> ubound x"
- by (simp add: ubound_def)
-
-lemma float_mult:
- "float (a1, e1) * float (a2, e2) =
- (float (a1 * a2, e1 + e2))"
- by (simp add: float_def pow2_add)
-
-lemma float_minus:
- "- (float (a,b)) = float (-a, b)"
- by (simp add: float_def)
-
-lemma zero_less_pow2:
- "0 < pow2 x"
-proof -
- {
- fix y
- have "0 <= y \<Longrightarrow> 0 < pow2 y"
- by (induct y, induct_tac n, simp_all add: pow2_add)
- }
- note helper=this
- show ?thesis
- apply (case_tac "0 <= x")
- apply (simp add: helper)
- apply (subst pow2_neg)
- apply (simp add: helper)
- done
-qed
-
-lemma zero_le_float:
- "(0 <= float (a,b)) = (0 <= a)"
- apply (auto simp add: float_def)
- apply (auto simp add: zero_le_mult_iff zero_less_pow2)
- apply (insert zero_less_pow2[of b])
- apply (simp_all)
- done
-
-lemma float_le_zero:
- "(float (a,b) <= 0) = (a <= 0)"
- apply (auto simp add: float_def)
- apply (auto simp add: mult_le_0_iff)
- apply (insert zero_less_pow2[of b])
- apply auto
- done
-
-lemma float_abs:
- "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))"
- apply (auto simp add: abs_if)
- apply (simp_all add: zero_le_float[symmetric, of a b] float_minus)
- done
-
-lemma float_zero:
- "float (0, b) = 0"
- by (simp add: float_def)
-
-lemma float_pprt:
- "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))"
- by (auto simp add: zero_le_float float_le_zero float_zero)
-
-lemma pprt_lbound: "pprt (lbound x) = float (0, 0)"
- apply (simp add: float_def)
- apply (rule pprt_eq_0)
- apply (simp add: lbound_def)
- done
-
-lemma nprt_ubound: "nprt (ubound x) = float (0, 0)"
- apply (simp add: float_def)
- apply (rule nprt_eq_0)
- apply (simp add: ubound_def)
- done
-
-lemma float_nprt:
- "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))"
- by (auto simp add: zero_le_float float_le_zero float_zero)
-
-lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1"
- by auto
-
-lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)"
- by simp
-
-lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)"
- by simp
-
-lemma mult_left_one: "1 * a = (a::'a::semiring_1)"
- by simp
-
-lemma mult_right_one: "a * 1 = (a::'a::semiring_1)"
- by simp
-
-lemma int_pow_0: "(a::int)^(Numeral0) = 1"
- by simp
-
-lemma int_pow_1: "(a::int)^(Numeral1) = a"
- by simp
-
-lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0"
- by simp
-
-lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1"
- by simp
-
-lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0"
- by simp
-
-lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1"
- by simp
-
-lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1"
- by simp
-
-lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1"
-proof -
- have 1:"((-1)::nat) = 0"
- by simp
- show ?thesis by (simp add: 1)
-qed
-
-lemma fst_cong: "a=a' \<Longrightarrow> fst (a,b) = fst (a',b)"
- by simp
-
-lemma snd_cong: "b=b' \<Longrightarrow> snd (a,b) = snd (a,b')"
- by simp
-
-lemma lift_bool: "x \<Longrightarrow> x=True"
- by simp
-
-lemma nlift_bool: "~x \<Longrightarrow> x=False"
- by simp
-
-lemma not_false_eq_true: "(~ False) = True" by simp
-
-lemma not_true_eq_false: "(~ True) = False" by simp
-
-lemmas binarith =
- normalize_bin_simps
- pred_bin_simps succ_bin_simps
- add_bin_simps minus_bin_simps mult_bin_simps
-
-lemma int_eq_number_of_eq:
- "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)"
- by (rule eq_number_of_eq)
-
-lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)"
- by (simp only: iszero_number_of_Pls)
-
-lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))"
- by simp
-
-lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)"
- by simp
-
-lemma int_iszero_number_of_Bit1: "\<not> iszero ((number_of (Int.Bit1 w))::int)"
- by simp
-
-lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)"
- unfolding neg_def number_of_is_id by simp
-
-lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))"
- by simp
-
-lemma int_neg_number_of_Min: "neg (-1::int)"
- by simp
-
-lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)"
- by simp
-
-lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)"
- by simp
-
-lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (y + (uminus x)))::int))"
- unfolding neg_def number_of_is_id by (simp add: not_less)
-
-lemmas intarithrel =
- int_eq_number_of_eq
- lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0
- lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min]
- int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq
-
-lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)"
- by simp
-
-lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))"
- by simp
-
-lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)"
- by simp
-
-lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)"
- by simp
-
-lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym
-
-lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of
-
-lemmas powerarith = nat_number_of zpower_number_of_even
- zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
- zpower_Pls zpower_Min
-
-lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0
- float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound
-
-(* for use with the compute oracle *)
-lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false
-
-use "~~/src/HOL/Tools/float_arith.ML"
-
-end
--- a/src/HOL/Tools/ComputeHOL.thy Wed Sep 02 21:33:16 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,191 +0,0 @@
-theory ComputeHOL
-imports Complex_Main "~~/src/Tools/Compute_Oracle/Compute_Oracle"
-begin
-
-lemma Trueprop_eq_eq: "Trueprop X == (X == True)" by (simp add: atomize_eq)
-lemma meta_eq_trivial: "x == y \<Longrightarrow> x == y" by simp
-lemma meta_eq_imp_eq: "x == y \<Longrightarrow> x = y" by auto
-lemma eq_trivial: "x = y \<Longrightarrow> x = y" by auto
-lemma bool_to_true: "x :: bool \<Longrightarrow> x == True" by simp
-lemma transmeta_1: "x = y \<Longrightarrow> y == z \<Longrightarrow> x = z" by simp
-lemma transmeta_2: "x == y \<Longrightarrow> y = z \<Longrightarrow> x = z" by simp
-lemma transmeta_3: "x == y \<Longrightarrow> y == z \<Longrightarrow> x = z" by simp
-
-
-(**** compute_if ****)
-
-lemma If_True: "If True = (\<lambda> x y. x)" by ((rule ext)+,auto)
-lemma If_False: "If False = (\<lambda> x y. y)" by ((rule ext)+, auto)
-
-lemmas compute_if = If_True If_False
-
-(**** compute_bool ****)
-
-lemma bool1: "(\<not> True) = False" by blast
-lemma bool2: "(\<not> False) = True" by blast
-lemma bool3: "(P \<and> True) = P" by blast
-lemma bool4: "(True \<and> P) = P" by blast
-lemma bool5: "(P \<and> False) = False" by blast
-lemma bool6: "(False \<and> P) = False" by blast
-lemma bool7: "(P \<or> True) = True" by blast
-lemma bool8: "(True \<or> P) = True" by blast
-lemma bool9: "(P \<or> False) = P" by blast
-lemma bool10: "(False \<or> P) = P" by blast
-lemma bool11: "(True \<longrightarrow> P) = P" by blast
-lemma bool12: "(P \<longrightarrow> True) = True" by blast
-lemma bool13: "(True \<longrightarrow> P) = P" by blast
-lemma bool14: "(P \<longrightarrow> False) = (\<not> P)" by blast
-lemma bool15: "(False \<longrightarrow> P) = True" by blast
-lemma bool16: "(False = False) = True" by blast
-lemma bool17: "(True = True) = True" by blast
-lemma bool18: "(False = True) = False" by blast
-lemma bool19: "(True = False) = False" by blast
-
-lemmas compute_bool = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10 bool11 bool12 bool13 bool14 bool15 bool16 bool17 bool18 bool19
-
-
-(*** compute_pair ***)
-
-lemma compute_fst: "fst (x,y) = x" by simp
-lemma compute_snd: "snd (x,y) = y" by simp
-lemma compute_pair_eq: "((a, b) = (c, d)) = (a = c \<and> b = d)" by auto
-
-lemma prod_case_simp: "prod_case f (x,y) = f x y" by simp
-
-lemmas compute_pair = compute_fst compute_snd compute_pair_eq prod_case_simp
-
-(*** compute_option ***)
-
-lemma compute_the: "the (Some x) = x" by simp
-lemma compute_None_Some_eq: "(None = Some x) = False" by auto
-lemma compute_Some_None_eq: "(Some x = None) = False" by auto
-lemma compute_None_None_eq: "(None = None) = True" by auto
-lemma compute_Some_Some_eq: "(Some x = Some y) = (x = y)" by auto
-
-definition
- option_case_compute :: "'b option \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
-where
- "option_case_compute opt a f = option_case a f opt"
-
-lemma option_case_compute: "option_case = (\<lambda> a f opt. option_case_compute opt a f)"
- by (simp add: option_case_compute_def)
-
-lemma option_case_compute_None: "option_case_compute None = (\<lambda> a f. a)"
- apply (rule ext)+
- apply (simp add: option_case_compute_def)
- done
-
-lemma option_case_compute_Some: "option_case_compute (Some x) = (\<lambda> a f. f x)"
- apply (rule ext)+
- apply (simp add: option_case_compute_def)
- done
-
-lemmas compute_option_case = option_case_compute option_case_compute_None option_case_compute_Some
-
-lemmas compute_option = compute_the compute_None_Some_eq compute_Some_None_eq compute_None_None_eq compute_Some_Some_eq compute_option_case
-
-(**** compute_list_length ****)
-
-lemma length_cons:"length (x#xs) = 1 + (length xs)"
- by simp
-
-lemma length_nil: "length [] = 0"
- by simp
-
-lemmas compute_list_length = length_nil length_cons
-
-(*** compute_list_case ***)
-
-definition
- list_case_compute :: "'b list \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b list \<Rightarrow> 'a) \<Rightarrow> 'a"
-where
- "list_case_compute l a f = list_case a f l"
-
-lemma list_case_compute: "list_case = (\<lambda> (a::'a) f (l::'b list). list_case_compute l a f)"
- apply (rule ext)+
- apply (simp add: list_case_compute_def)
- done
-
-lemma list_case_compute_empty: "list_case_compute ([]::'b list) = (\<lambda> (a::'a) f. a)"
- apply (rule ext)+
- apply (simp add: list_case_compute_def)
- done
-
-lemma list_case_compute_cons: "list_case_compute (u#v) = (\<lambda> (a::'a) f. (f (u::'b) v))"
- apply (rule ext)+
- apply (simp add: list_case_compute_def)
- done
-
-lemmas compute_list_case = list_case_compute list_case_compute_empty list_case_compute_cons
-
-(*** compute_list_nth ***)
-(* Of course, you will need computation with nats for this to work \<dots> *)
-
-lemma compute_list_nth: "((x#xs) ! n) = (if n = 0 then x else (xs ! (n - 1)))"
- by (cases n, auto)
-
-(*** compute_list ***)
-
-lemmas compute_list = compute_list_case compute_list_length compute_list_nth
-
-(*** compute_let ***)
-
-lemmas compute_let = Let_def
-
-(***********************)
-(* Everything together *)
-(***********************)
-
-lemmas compute_hol = compute_if compute_bool compute_pair compute_option compute_list compute_let
-
-ML {*
-signature ComputeHOL =
-sig
- val prep_thms : thm list -> thm list
- val to_meta_eq : thm -> thm
- val to_hol_eq : thm -> thm
- val symmetric : thm -> thm
- val trans : thm -> thm -> thm
-end
-
-structure ComputeHOL : ComputeHOL =
-struct
-
-local
-fun lhs_of eq = fst (Thm.dest_equals (cprop_of eq));
-in
-fun rewrite_conv [] ct = raise CTERM ("rewrite_conv", [])
- | rewrite_conv (eq :: eqs) ct =
- Thm.instantiate (Thm.match (lhs_of eq, ct)) eq
- handle Pattern.MATCH => rewrite_conv eqs ct;
-end
-
-val convert_conditions = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.try_conv (rewrite_conv [@{thm "Trueprop_eq_eq"}])))
-
-val eq_th = @{thm "HOL.eq_reflection"}
-val meta_eq_trivial = @{thm "ComputeHOL.meta_eq_trivial"}
-val bool_to_true = @{thm "ComputeHOL.bool_to_true"}
-
-fun to_meta_eq th = eq_th OF [th] handle THM _ => meta_eq_trivial OF [th] handle THM _ => bool_to_true OF [th]
-
-fun to_hol_eq th = @{thm "meta_eq_imp_eq"} OF [th] handle THM _ => @{thm "eq_trivial"} OF [th]
-
-fun prep_thms ths = map (convert_conditions o to_meta_eq) ths
-
-fun symmetric th = @{thm "HOL.sym"} OF [th] handle THM _ => @{thm "Pure.symmetric"} OF [th]
-
-local
- val trans_HOL = @{thm "HOL.trans"}
- val trans_HOL_1 = @{thm "ComputeHOL.transmeta_1"}
- val trans_HOL_2 = @{thm "ComputeHOL.transmeta_2"}
- val trans_HOL_3 = @{thm "ComputeHOL.transmeta_3"}
- fun tr [] th1 th2 = trans_HOL OF [th1, th2]
- | tr (t::ts) th1 th2 = (t OF [th1, th2] handle THM _ => tr ts th1 th2)
-in
- fun trans th1 th2 = tr [trans_HOL, trans_HOL_1, trans_HOL_2, trans_HOL_3] th1 th2
-end
-
-end
-*}
-
-end
--- a/src/HOL/Tools/ComputeNumeral.thy Wed Sep 02 21:33:16 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,195 +0,0 @@
-theory ComputeNumeral
-imports ComputeHOL ComputeFloat
-begin
-
-(* normalization of bit strings *)
-lemmas bitnorm = normalize_bin_simps
-
-(* neg for bit strings *)
-lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def)
-lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto
-lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto
-lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto
-lemmas bitneg = neg1 neg2 neg3 neg4
-
-(* iszero for bit strings *)
-lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def)
-lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp
-lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto
-lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+ apply simp by arith
-lemmas bitiszero = iszero1 iszero2 iszero3 iszero4
-
-(* lezero for bit strings *)
-constdefs
- "lezero x == (x \<le> 0)"
-lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto
-lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto
-lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto
-lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto
-lemmas bitlezero = lezero1 lezero2 lezero3 lezero4
-
-(* equality for bit strings *)
-lemmas biteq = eq_bin_simps
-
-(* x < y for bit strings *)
-lemmas bitless = less_bin_simps
-
-(* x \<le> y for bit strings *)
-lemmas bitle = le_bin_simps
-
-(* succ for bit strings *)
-lemmas bitsucc = succ_bin_simps
-
-(* pred for bit strings *)
-lemmas bitpred = pred_bin_simps
-
-(* unary minus for bit strings *)
-lemmas bituminus = minus_bin_simps
-
-(* addition for bit strings *)
-lemmas bitadd = add_bin_simps
-
-(* multiplication for bit strings *)
-lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def)
-lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute, simp add: mult_Min)
-lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0)
-lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp
-lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)"
- unfolding Bit0_def Bit1_def by (simp add: algebra_simps)
-lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1
-
-lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul
-
-constdefs
- "nat_norm_number_of (x::nat) == x"
-
-lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)"
- apply (simp add: nat_norm_number_of_def)
- unfolding lezero_def iszero_def neg_def
- apply (simp add: numeral_simps)
- done
-
-(* Normalization of nat literals *)
-lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto
-lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)" by auto
-lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of
-
-(* Suc *)
-lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id)
-
-(* Addition for nat *)
-lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))"
- unfolding nat_number_of_def number_of_is_id neg_def
- by auto
-
-(* Subtraction for nat *)
-lemma natsub: "(number_of x) - ((number_of y)::nat) =
- (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))"
- unfolding nat_norm_number_of
- by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def)
-
-(* Multiplication for nat *)
-lemma natmul: "(number_of x) * ((number_of y)::nat) =
- (if neg x then 0 else (if neg y then 0 else number_of (x * y)))"
- unfolding nat_number_of_def number_of_is_id neg_def
- by (simp add: nat_mult_distrib)
-
-lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \<and> lezero y) \<or> (x = y))"
- by (auto simp add: iszero_def lezero_def neg_def number_of_is_id)
-
-lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \<and> (\<not> (lezero y)))"
- by (simp add: lezero_def numeral_simps not_le)
-
-lemma natle: "(((number_of x)::nat) \<le> (number_of y)) = (y < x \<longrightarrow> lezero x)"
- by (auto simp add: number_of_is_id lezero_def nat_number_of_def)
-
-fun natfac :: "nat \<Rightarrow> nat"
-where
- "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))"
-
-lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps
-
-lemma number_eq: "(((number_of x)::'a::{number_ring, ordered_idom}) = (number_of y)) = (x = y)"
- unfolding number_of_eq
- apply simp
- done
-
-lemma number_le: "(((number_of x)::'a::{number_ring, ordered_idom}) \<le> (number_of y)) = (x \<le> y)"
- unfolding number_of_eq
- apply simp
- done
-
-lemma number_less: "(((number_of x)::'a::{number_ring, ordered_idom}) < (number_of y)) = (x < y)"
- unfolding number_of_eq
- apply simp
- done
-
-lemma number_diff: "((number_of x)::'a::{number_ring, ordered_idom}) - number_of y = number_of (x + (- y))"
- apply (subst diff_number_of_eq)
- apply simp
- done
-
-lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric]
-
-lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less
-
-lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)"
- by (simp only: real_of_nat_number_of number_of_is_id)
-
-lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)"
- by simp
-
-lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of
-
-lemmas zpowerarith = zpower_number_of_even
- zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
- zpower_Pls zpower_Min
-
-(* div, mod *)
-
-lemma adjust: "adjust b (q, r) = (if 0 \<le> r - b then (2 * q + 1, r - b) else (2 * q, r))"
- by (auto simp only: adjust_def)
-
-lemma negateSnd: "negateSnd (q, r) = (q, -r)"
- by (simp add: negateSnd_def)
-
-lemma divmod: "IntDiv.divmod a b = (if 0\<le>a then
- if 0\<le>b then posDivAlg a b
- else if a=0 then (0, 0)
- else negateSnd (negDivAlg (-a) (-b))
- else
- if 0<b then negDivAlg a b
- else negateSnd (posDivAlg (-a) (-b)))"
- by (auto simp only: IntDiv.divmod_def)
-
-lemmas compute_div_mod = div_def mod_def divmod adjust negateSnd posDivAlg.simps negDivAlg.simps
-
-
-
-(* collecting all the theorems *)
-
-lemma even_Pls: "even (Int.Pls) = True"
- apply (unfold Pls_def even_def)
- by simp
-
-lemma even_Min: "even (Int.Min) = False"
- apply (unfold Min_def even_def)
- by simp
-
-lemma even_B0: "even (Int.Bit0 x) = True"
- apply (unfold Bit0_def)
- by simp
-
-lemma even_B1: "even (Int.Bit1 x) = False"
- apply (unfold Bit1_def)
- by simp
-
-lemma even_number_of: "even ((number_of w)::int) = even w"
- by (simp only: number_of_is_id)
-
-lemmas compute_even = even_Pls even_Min even_B0 even_B1 even_number_of
-
-lemmas compute_numeral = compute_if compute_let compute_pair compute_bool
- compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even
-
-end
--- a/src/Pure/General/swing_thread.scala Wed Sep 02 21:33:16 2009 +0200
+++ b/src/Pure/General/swing_thread.scala Wed Sep 02 21:34:13 2009 +0200
@@ -13,17 +13,23 @@
object Swing_Thread
{
+ /* checks */
+
+ def assert() = Predef.assert(SwingUtilities.isEventDispatchThread())
+ def require() = Predef.require(SwingUtilities.isEventDispatchThread())
+
+
/* main dispatch queue */
def now[A](body: => A): A = {
var result: Option[A] = None
- if (SwingUtilities.isEventDispatchThread) { result = Some(body) }
+ if (SwingUtilities.isEventDispatchThread()) { result = Some(body) }
else SwingUtilities.invokeAndWait(new Runnable { def run = { result = Some(body) } })
result.get
}
def later(body: => Unit) {
- if (SwingUtilities.isEventDispatchThread) body
+ if (SwingUtilities.isEventDispatchThread()) body
else SwingUtilities.invokeLater(new Runnable { def run = body })
}
--- a/src/Pure/ML/ml_compiler_polyml-5.3.ML Wed Sep 02 21:33:16 2009 +0200
+++ b/src/Pure/ML/ml_compiler_polyml-5.3.ML Wed Sep 02 21:34:13 2009 +0200
@@ -1,7 +1,7 @@
(* Title: Pure/ML/ml_compiler_polyml-5.3.ML
Author: Makarius
-Advanced runtime compilation for Poly/ML 5.3 (SVN 773).
+Advanced runtime compilation for Poly/ML 5.3 (SVN 839).
*)
signature ML_COMPILER =
@@ -74,8 +74,8 @@
(* input *)
val location_props =
- Markup.markup (Markup.position |> Markup.properties
- (filter (member (op =) [Markup.idN, Markup.fileN] o #1) (Position.properties_of pos))) "";
+ op ^ (YXML.output_markup (Markup.position |> Markup.properties
+ (filter (member (op =) [Markup.idN, Markup.fileN] o #1) (Position.properties_of pos))));
val input = toks |> maps (fn tok =>
let