merged
authorhuffman
Fri, 20 Feb 2009 16:07:20 -0800
changeset 30033 e54d4d41fe8f
parent 30032 c7f0c1b8001b (diff)
parent 30031 bd786c37af84 (current diff)
child 30035 7f4d475a6c50
child 30036 3a074e3a9a18
merged
--- a/src/HOL/Library/Numeral_Type.thy	Fri Feb 20 23:46:03 2009 +0100
+++ b/src/HOL/Library/Numeral_Type.thy	Fri Feb 20 16:07:20 2009 -0800
@@ -267,6 +267,22 @@
 
 subsection {* Number ring instances *}
 
+text {*
+  Unfortunately a number ring instance is not possible for
+  @{typ num1}, since 0 and 1 are not distinct.
+*}
+
+instantiation num1 :: "{comm_ring,comm_monoid_mult,number,recpower}"
+begin
+
+lemma num1_eq_iff: "(x::num1) = (y::num1) \<longleftrightarrow> True"
+  by (induct x, induct y) simp
+
+instance proof
+qed (simp_all add: num1_eq_iff)
+
+end
+
 instantiation
   bit0 and bit1 :: (finite) "{zero,one,plus,times,uminus,minus,power}"
 begin