number_ring instances for numeral types
authorhuffman
Thu, 19 Feb 2009 16:51:46 -0800
changeset 29997 f6756c097c2d
parent 29996 c09f348ca88a
child 29998 19e1ef628b25
number_ring instances for numeral types
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) \<notin> 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 \<in> {0 ..< 2 * int CARD('a)}"
+    by simp
+qed
+
+typedef (open) 'a bit1 = "{0 ..< 1 + 2 * int CARD('a::finite)}"
+proof
+  show "0 \<in> {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} \<Rightarrow> int"
+  and Abs :: "int \<Rightarrow> 'a::{zero,one,plus,times,uminus,minus,power}"
+  assumes type: "type_definition Rep Abs {0..<n}"
+  and size1: "1 < n"
+  and zero_def: "0 = Abs 0"
+  and one_def:  "1 = Abs 1"
+  and add_def:  "x + y = Abs ((Rep x + Rep y) mod n)"
+  and mult_def: "x * y = Abs ((Rep x * Rep y) mod n)"
+  and diff_def: "x - y = Abs ((Rep x - Rep y) mod n)"
+  and minus_def: "- x = Abs ((- Rep x) mod n)"
+  and power_def: "x ^ k = Abs (Rep x ^ k mod n)"
+begin
+
+lemma size0: "0 < n"
+by (cut_tac size1, simp)
+
+lemmas definitions =
+  zero_def one_def add_def mult_def minus_def diff_def power_def
+
+lemma Rep_less_n: "Rep x < n"
+by (rule type_definition.Rep [OF type, simplified, THEN conjunct2])
+
+lemma Rep_le_n: "Rep x \<le> n"
+by (rule Rep_less_n [THEN order_less_imp_le])
+
+lemma Rep_inject_sym: "x = y \<longleftrightarrow> 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 \<in> {0..<n} \<Longrightarrow> 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} \<Rightarrow> int"
+  and Abs :: "int \<Rightarrow> '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) \<longleftrightarrow> 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: "\<And>z. \<lbrakk>(x::'a) = of_int z; 0 \<le> z; z < n\<rbrakk> \<Longrightarrow> 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:
+  "(\<And>z. \<lbrakk>0 \<le> z; z < n\<rbrakk> \<Longrightarrow> P (of_int z)) \<Longrightarrow> 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 \<Rightarrow> 'a bit0" where
+  "Abs_bit0' x = Abs_bit0 (x mod (2 * int CARD('a)))"
+
+definition Abs_bit1' :: "int \<Rightarrow> '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 \<Rightarrow> int"
+           "Abs_bit0 :: int \<Rightarrow> '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 \<Rightarrow> int"
+           "Abs_bit1 :: int \<Rightarrow> '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 \<Rightarrow> int"
+           "Abs_bit0 :: int \<Rightarrow> 'a::finite bit0"
+  ..
+
+interpretation bit1!:
+  mod_ring "1 + 2 * int CARD('a::finite)"
+           "Rep_bit1 :: 'a::finite bit1 \<Rightarrow> int"
+           "Abs_bit1 :: int \<Rightarrow> '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