diff -r c4202c67fe3e -r 6da615cef733 src/HOL/Int.thy --- a/src/HOL/Int.thy Wed Apr 02 12:32:53 2008 +0200 +++ b/src/HOL/Int.thy Wed Apr 02 15:58:26 2008 +0200 @@ -1797,9 +1797,125 @@ subsection {* Configuration of the code generator *} -instance int :: eq .. +code_datatype Pls Min Bit0 Bit1 "number_of \ int \ int" + +lemmas pred_succ_numeral_code [code func] = + pred_bin_simps succ_bin_simps + +lemmas plus_numeral_code [code func] = + add_bin_simps + arith_extra_simps(1) [where 'a = int] + +lemmas minus_numeral_code [code func] = + minus_bin_simps + arith_extra_simps(2) [where 'a = int] + arith_extra_simps(5) [where 'a = int] + +lemmas times_numeral_code [code func] = + mult_bin_simps + arith_extra_simps(4) [where 'a = int] + +instantiation int :: eq +begin + +definition [code func del]: "eq k l \ k - l = (0\int)" + +instance by default (simp add: eq_int_def) + +end + +lemma eq_number_of_int_code [code func]: + "eq (number_of k \ int) (number_of l) \ eq k l" + unfolding eq_int_def number_of_is_id .. + +lemma eq_int_code [code func]: + "eq Int.Pls Int.Pls \ True" + "eq Int.Pls Int.Min \ False" + "eq Int.Pls (Int.Bit0 k2) \ eq Int.Pls k2" + "eq Int.Pls (Int.Bit1 k2) \ False" + "eq Int.Min Int.Pls \ False" + "eq Int.Min Int.Min \ True" + "eq Int.Min (Int.Bit0 k2) \ False" + "eq Int.Min (Int.Bit1 k2) \ eq Int.Min k2" + "eq (Int.Bit0 k1) Int.Pls \ eq Int.Pls k1" + "eq (Int.Bit1 k1) Int.Pls \ False" + "eq (Int.Bit0 k1) Int.Min \ False" + "eq (Int.Bit1 k1) Int.Min \ eq Int.Min k1" + "eq (Int.Bit0 k1) (Int.Bit0 k2) \ eq k1 k2" + "eq (Int.Bit0 k1) (Int.Bit1 k2) \ False" + "eq (Int.Bit1 k1) (Int.Bit0 k2) \ False" + "eq (Int.Bit1 k1) (Int.Bit1 k2) \ eq k1 k2" + unfolding eq_number_of_int_code [symmetric, of Int.Pls] + eq_number_of_int_code [symmetric, of Int.Min] + eq_number_of_int_code [symmetric, of "Int.Bit0 k1"] + eq_number_of_int_code [symmetric, of "Int.Bit1 k1"] + eq_number_of_int_code [symmetric, of "Int.Bit0 k2"] + eq_number_of_int_code [symmetric, of "Int.Bit1 k2"] + by (simp_all add: eq Pls_def, + simp_all only: Min_def succ_def pred_def number_of_is_id) + (auto simp add: iszero_def) -code_datatype Pls Min Bit0 Bit1 "number_of \ int \ int" +lemma less_eq_number_of_int_code [code func]: + "(number_of k \ int) \ number_of l \ k \ l" + unfolding number_of_is_id .. + +lemma less_eq_int_code [code func]: + "Int.Pls \ Int.Pls \ True" + "Int.Pls \ Int.Min \ False" + "Int.Pls \ Int.Bit0 k \ Int.Pls \ k" + "Int.Pls \ Int.Bit1 k \ Int.Pls \ k" + "Int.Min \ Int.Pls \ True" + "Int.Min \ Int.Min \ True" + "Int.Min \ Int.Bit0 k \ Int.Min < k" + "Int.Min \ Int.Bit1 k \ Int.Min \ k" + "Int.Bit0 k \ Int.Pls \ k \ Int.Pls" + "Int.Bit1 k \ Int.Pls \ k < Int.Pls" + "Int.Bit0 k \ Int.Min \ k \ Int.Min" + "Int.Bit1 k \ Int.Min \ k \ Int.Min" + "Int.Bit0 k1 \ Int.Bit0 k2 \ k1 \ k2" + "Int.Bit0 k1 \ Int.Bit1 k2 \ k1 \ k2" + "Int.Bit1 k1 \ Int.Bit0 k2 \ k1 < k2" + "Int.Bit1 k1 \ Int.Bit1 k2 \ k1 \ k2" + unfolding less_eq_number_of_int_code [symmetric, of Int.Pls] + less_eq_number_of_int_code [symmetric, of Int.Min] + less_eq_number_of_int_code [symmetric, of "Int.Bit0 k1"] + less_eq_number_of_int_code [symmetric, of "Int.Bit1 k1"] + less_eq_number_of_int_code [symmetric, of "Int.Bit0 k2"] + less_eq_number_of_int_code [symmetric, of "Int.Bit1 k2"] + by (simp_all add: Pls_def, simp_all only: Min_def succ_def pred_def number_of_is_id) + (auto simp add: neg_def linorder_not_less group_simps + zle_add1_eq_le [symmetric] del: iffI , auto simp only: Bit0_def Bit1_def) + +lemma less_number_of_int_code [code func]: + "(number_of k \ int) < number_of l \ k < l" + unfolding number_of_is_id .. + +lemma less_int_code [code func]: + "Int.Pls < Int.Pls \ False" + "Int.Pls < Int.Min \ False" + "Int.Pls < Int.Bit0 k \ Int.Pls < k" + "Int.Pls < Int.Bit1 k \ Int.Pls \ k" + "Int.Min < Int.Pls \ True" + "Int.Min < Int.Min \ False" + "Int.Min < Int.Bit0 k \ Int.Min < k" + "Int.Min < Int.Bit1 k \ Int.Min < k" + "Int.Bit0 k < Int.Pls \ k < Int.Pls" + "Int.Bit1 k < Int.Pls \ k < Int.Pls" + "Int.Bit0 k < Int.Min \ k \ Int.Min" + "Int.Bit1 k < Int.Min \ k < Int.Min" + "Int.Bit0 k1 < Int.Bit0 k2 \ k1 < k2" + "Int.Bit0 k1 < Int.Bit1 k2 \ k1 \ k2" + "Int.Bit1 k1 < Int.Bit0 k2 \ k1 < k2" + "Int.Bit1 k1 < Int.Bit1 k2 \ k1 < k2" + unfolding less_number_of_int_code [symmetric, of Int.Pls] + less_number_of_int_code [symmetric, of Int.Min] + less_number_of_int_code [symmetric, of "Int.Bit0 k1"] + less_number_of_int_code [symmetric, of "Int.Bit1 k1"] + less_number_of_int_code [symmetric, of "Int.Bit0 k2"] + less_number_of_int_code [symmetric, of "Int.Bit1 k2"] + by (simp_all add: Pls_def, simp_all only: Min_def succ_def pred_def number_of_is_id) + (auto simp add: neg_def group_simps zle_add1_eq_le [symmetric] del: iffI, + auto simp only: Bit0_def Bit1_def) definition int_aux :: "nat \ int \ int" where