src/HOL/Library/Numeral_Type.thy
changeset 62348 9a5f43dac883
parent 61649 268d88ec9087
child 64593 50c715579715
equal deleted inserted replaced
62347:2230b7047376 62348:9a5f43dac883
   228 interpretation bit0:
   228 interpretation bit0:
   229   mod_type "int CARD('a::finite bit0)"
   229   mod_type "int CARD('a::finite bit0)"
   230            "Rep_bit0 :: 'a::finite bit0 \<Rightarrow> int"
   230            "Rep_bit0 :: 'a::finite bit0 \<Rightarrow> int"
   231            "Abs_bit0 :: int \<Rightarrow> 'a::finite bit0"
   231            "Abs_bit0 :: int \<Rightarrow> 'a::finite bit0"
   232 apply (rule mod_type.intro)
   232 apply (rule mod_type.intro)
   233 apply (simp add: int_mult type_definition_bit0)
   233 apply (simp add: of_nat_mult type_definition_bit0)
   234 apply (rule one_less_int_card)
   234 apply (rule one_less_int_card)
   235 apply (rule zero_bit0_def)
   235 apply (rule zero_bit0_def)
   236 apply (rule one_bit0_def)
   236 apply (rule one_bit0_def)
   237 apply (rule plus_bit0_def [unfolded Abs_bit0'_def])
   237 apply (rule plus_bit0_def [unfolded Abs_bit0'_def])
   238 apply (rule times_bit0_def [unfolded Abs_bit0'_def])
   238 apply (rule times_bit0_def [unfolded Abs_bit0'_def])
   243 interpretation bit1:
   243 interpretation bit1:
   244   mod_type "int CARD('a::finite bit1)"
   244   mod_type "int CARD('a::finite bit1)"
   245            "Rep_bit1 :: 'a::finite bit1 \<Rightarrow> int"
   245            "Rep_bit1 :: 'a::finite bit1 \<Rightarrow> int"
   246            "Abs_bit1 :: int \<Rightarrow> 'a::finite bit1"
   246            "Abs_bit1 :: int \<Rightarrow> 'a::finite bit1"
   247 apply (rule mod_type.intro)
   247 apply (rule mod_type.intro)
   248 apply (simp add: int_mult type_definition_bit1)
   248 apply (simp add: of_nat_mult type_definition_bit1)
   249 apply (rule one_less_int_card)
   249 apply (rule one_less_int_card)
   250 apply (rule zero_bit1_def)
   250 apply (rule zero_bit1_def)
   251 apply (rule one_bit1_def)
   251 apply (rule one_bit1_def)
   252 apply (rule plus_bit1_def [unfolded Abs_bit1'_def])
   252 apply (rule plus_bit1_def [unfolded Abs_bit1'_def])
   253 apply (rule times_bit1_def [unfolded Abs_bit1'_def])
   253 apply (rule times_bit1_def [unfolded Abs_bit1'_def])