# HG changeset patch # User huffman # Date 1235091106 28800 # Node ID f6756c097c2d3b0bb60f35fe7a7edbf3d4ae5c81 # Parent c09f348ca88a36c00d35554eb1de5fe86f635756 number_ring instances for numeral types diff -r c09f348ca88a -r f6756c097c2d src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy Thu Feb 19 12:37:03 2009 -0800 +++ b/src/HOL/Library/Numeral_Type.thy Thu Feb 19 16:51:46 2009 -0800 @@ -57,7 +57,7 @@ lemma card_option: "CARD('a::finite option) = Suc CARD('a)" unfolding insert_None_conv_UNIV [symmetric] apply (subgoal_tac "(None::'a option) \ range Some") - apply (simp add: finite card_image) + apply (simp add: card_image) apply fast done @@ -65,13 +65,26 @@ unfolding Pow_UNIV [symmetric] by (simp only: card_Pow finite numeral_2_eq_2) +lemma card_finite_pos [simp]: "0 < CARD('a::finite)" + unfolding neq0_conv [symmetric] by simp + subsection {* Numeral Types *} typedef (open) num0 = "UNIV :: nat set" .. typedef (open) num1 = "UNIV :: unit set" .. -typedef (open) 'a bit0 = "UNIV :: (bool * 'a) set" .. -typedef (open) 'a bit1 = "UNIV :: (bool * 'a) option set" .. + +typedef (open) 'a bit0 = "{0 ..< 2 * int CARD('a::finite)}" +proof + show "0 \ {0 ..< 2 * int CARD('a)}" + by simp +qed + +typedef (open) 'a bit1 = "{0 ..< 1 + 2 * int CARD('a::finite)}" +proof + show "0 \ {0 ..< 1 + 2 * int CARD('a)}" + by simp +qed instance num1 :: finite proof @@ -84,14 +97,14 @@ proof show "finite (UNIV::'a bit0 set)" unfolding type_definition.univ [OF type_definition_bit0] - using finite by (rule finite_imageI) + by simp qed instance bit1 :: (finite) finite proof show "finite (UNIV::'a bit1 set)" unfolding type_definition.univ [OF type_definition_bit1] - using finite by (rule finite_imageI) + by simp qed lemma card_num1: "CARD(num1) = 1" @@ -100,11 +113,11 @@ lemma card_bit0: "CARD('a::finite bit0) = 2 * CARD('a)" unfolding type_definition.card [OF type_definition_bit0] - by (simp only: card_prod card_bool) + by simp lemma card_bit1: "CARD('a::finite bit1) = Suc (2 * CARD('a))" unfolding type_definition.card [OF type_definition_bit1] - by (simp only: card_prod card_option card_bool) + by simp lemma card_num0: "CARD (num0) = 0" by (simp add: infinite_UNIV_nat card_eq_0_iff type_definition.card [OF type_definition_num0]) @@ -122,6 +135,230 @@ card_num0 +subsection {* Locale for modular arithmetic subtypes *} + +locale mod_type = + fixes n :: int + and Rep :: "'a::{zero,one,plus,times,uminus,minus,power} \ int" + and Abs :: "int \ 'a::{zero,one,plus,times,uminus,minus,power}" + assumes type: "type_definition Rep Abs {0.. n" +by (rule Rep_less_n [THEN order_less_imp_le]) + +lemma Rep_inject_sym: "x = y \ Rep x = Rep y" +by (rule type_definition.Rep_inject [OF type, symmetric]) + +lemma Rep_inverse: "Abs (Rep x) = x" +by (rule type_definition.Rep_inverse [OF type]) + +lemma Abs_inverse: "m \ {0.. Rep (Abs m) = m" +by (rule type_definition.Abs_inverse [OF type]) + +lemma Rep_Abs_mod: "Rep (Abs (m mod n)) = m mod n" +by (simp add: Abs_inverse IntDiv.pos_mod_conj [OF size0]) + +lemma Rep_Abs_0: "Rep (Abs 0) = 0" +by (simp add: Abs_inverse size0) + +lemma Rep_0: "Rep 0 = 0" +by (simp add: zero_def Rep_Abs_0) + +lemma Rep_Abs_1: "Rep (Abs 1) = 1" +by (simp add: Abs_inverse size1) + +lemma Rep_1: "Rep 1 = 1" +by (simp add: one_def Rep_Abs_1) + +lemma Rep_mod: "Rep x mod n = Rep x" +apply (rule_tac x=x in type_definition.Abs_cases [OF type]) +apply (simp add: type_definition.Abs_inverse [OF type]) +apply (simp add: mod_pos_pos_trivial) +done + +lemmas Rep_simps = + Rep_inject_sym Rep_inverse Rep_Abs_mod Rep_mod Rep_Abs_0 Rep_Abs_1 + +lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)" +apply (intro_classes, unfold definitions) +apply (simp_all add: Rep_simps zmod_simps ring_simps) +done + +lemma recpower: "OFCLASS('a, recpower_class)" +apply (intro_classes, unfold definitions) +apply (simp_all add: Rep_simps zmod_simps add_ac mult_assoc + mod_pos_pos_trivial size1) +done + +end + +locale mod_ring = mod_type + + constrains n :: int + and Rep :: "'a::{number_ring,power} \ int" + and Abs :: "int \ 'a::{number_ring,power}" +begin + +lemma of_nat_eq: "of_nat k = Abs (int k mod n)" +apply (induct k) +apply (simp add: zero_def) +apply (simp add: Rep_simps add_def one_def zmod_simps add_ac) +done + +lemma of_int_eq: "of_int z = Abs (z mod n)" +apply (cases z rule: int_diff_cases) +apply (simp add: Rep_simps of_nat_eq diff_def zmod_simps) +done + +lemma Rep_number_of: + "Rep (number_of w) = number_of w mod n" +by (simp add: number_of_eq of_int_eq Rep_Abs_mod) + +lemma iszero_number_of: + "iszero (number_of w::'a) \ number_of w mod n = 0" +by (simp add: Rep_simps number_of_eq of_int_eq iszero_def zero_def) + +lemma cases: + assumes 1: "\z. \(x::'a) = of_int z; 0 \ z; z < n\ \ P" + shows "P" +apply (cases x rule: type_definition.Abs_cases [OF type]) +apply (rule_tac z="y" in 1) +apply (simp_all add: of_int_eq mod_pos_pos_trivial) +done + +lemma induct: + "(\z. \0 \ z; z < n\ \ P (of_int z)) \ P (x::'a)" +by (cases x rule: cases) simp + +end + + +subsection {* Number ring instances *} + +instantiation + bit0 and bit1 :: (finite) "{zero,one,plus,times,uminus,minus,power}" +begin + +definition Abs_bit0' :: "int \ 'a bit0" where + "Abs_bit0' x = Abs_bit0 (x mod (2 * int CARD('a)))" + +definition Abs_bit1' :: "int \ 'a bit1" where + "Abs_bit1' x = Abs_bit1 (x mod (1 + 2 * int CARD('a)))" + +definition "0 = Abs_bit0 0" +definition "1 = Abs_bit0 1" +definition "x + y = Abs_bit0' (Rep_bit0 x + Rep_bit0 y)" +definition "x * y = Abs_bit0' (Rep_bit0 x * Rep_bit0 y)" +definition "x - y = Abs_bit0' (Rep_bit0 x - Rep_bit0 y)" +definition "- x = Abs_bit0' (- Rep_bit0 x)" +definition "x ^ k = Abs_bit0' (Rep_bit0 x ^ k)" + +definition "0 = Abs_bit1 0" +definition "1 = Abs_bit1 1" +definition "x + y = Abs_bit1' (Rep_bit1 x + Rep_bit1 y)" +definition "x * y = Abs_bit1' (Rep_bit1 x * Rep_bit1 y)" +definition "x - y = Abs_bit1' (Rep_bit1 x - Rep_bit1 y)" +definition "- x = Abs_bit1' (- Rep_bit1 x)" +definition "x ^ k = Abs_bit1' (Rep_bit1 x ^ k)" + +instance .. + +end + +interpretation bit0!: + mod_type "2 * int CARD('a::finite)" + "Rep_bit0 :: 'a::finite bit0 \ int" + "Abs_bit0 :: int \ 'a::finite bit0" +apply (rule mod_type.intro) +apply (rule type_definition_bit0) +using card_finite_pos [where ?'a='a] apply arith +apply (rule zero_bit0_def) +apply (rule one_bit0_def) +apply (rule plus_bit0_def [unfolded Abs_bit0'_def]) +apply (rule times_bit0_def [unfolded Abs_bit0'_def]) +apply (rule minus_bit0_def [unfolded Abs_bit0'_def]) +apply (rule uminus_bit0_def [unfolded Abs_bit0'_def]) +apply (rule power_bit0_def [unfolded Abs_bit0'_def]) +done + +interpretation bit1!: + mod_type "1 + 2 * int CARD('a::finite)" + "Rep_bit1 :: 'a::finite bit1 \ int" + "Abs_bit1 :: int \ 'a::finite bit1" +apply (rule mod_type.intro) +apply (rule type_definition_bit1) +apply simp +apply (rule zero_bit1_def) +apply (rule one_bit1_def) +apply (rule plus_bit1_def [unfolded Abs_bit1'_def]) +apply (rule times_bit1_def [unfolded Abs_bit1'_def]) +apply (rule minus_bit1_def [unfolded Abs_bit1'_def]) +apply (rule uminus_bit1_def [unfolded Abs_bit1'_def]) +apply (rule power_bit1_def [unfolded Abs_bit1'_def]) +done + +instance bit0 :: (finite) "{comm_ring_1,recpower}" + by (rule bit0.comm_ring_1 bit0.recpower)+ + +instance bit1 :: (finite) "{comm_ring_1,recpower}" + by (rule bit1.comm_ring_1 bit1.recpower)+ + +instantiation bit0 and bit1 :: (finite) number_ring +begin + +definition "(number_of w :: _ bit0) = of_int w" + +definition "(number_of w :: _ bit1) = of_int w" + +instance proof +qed (rule number_of_bit0_def number_of_bit1_def)+ + +end + +interpretation bit0!: + mod_ring "2 * int CARD('a::finite)" + "Rep_bit0 :: 'a::finite bit0 \ int" + "Abs_bit0 :: int \ 'a::finite bit0" + .. + +interpretation bit1!: + mod_ring "1 + 2 * int CARD('a::finite)" + "Rep_bit1 :: 'a::finite bit1 \ int" + "Abs_bit1 :: int \ 'a::finite bit1" + .. + +text {* Set up cases, induction, and arithmetic *} + +lemmas bit0_cases [cases type: bit0, case_names of_int] = bit0.cases +lemmas bit1_cases [cases type: bit1, case_names of_int] = bit1.cases + +lemmas bit0_induct [induct type: bit0, case_names of_int] = bit0.induct +lemmas bit1_induct [induct type: bit1, case_names of_int] = bit1.induct + +lemmas bit0_iszero_number_of [simp] = bit0.iszero_number_of +lemmas bit1_iszero_number_of [simp] = bit1.iszero_number_of + +declare power_Suc [where ?'a="'a::finite bit0", standard, simp] +declare power_Suc [where ?'a="'a::finite bit1", standard, simp] + + subsection {* Syntax *} syntax @@ -221,5 +458,6 @@ lemma "CARD(0) = 0" by simp lemma "CARD(17) = 17" by simp +lemma "8 * 11 ^ 3 - 6 = (2::5)" by simp end