fixed and tuned
authorhaftmann
Fri, 25 Jan 2008 14:54:46 +0100
changeset 25967 dd602eb20f3f
parent 25966 74f6817870f9
child 25968 66cfe1d00be0
fixed and tuned
src/HOL/Library/Code_Index.thy
src/HOL/Library/Efficient_Nat.thy
src/HOL/Tools/numeral.ML
--- a/src/HOL/Library/Code_Index.thy	Fri Jan 25 14:54:44 2008 +0100
+++ b/src/HOL/Library/Code_Index.thy	Fri Jan 25 14:54:46 2008 +0100
@@ -17,6 +17,10 @@
 
 datatype index = index_of_nat nat
 
+lemma [code func]:
+  "index_size k = 0"
+  by (cases k) simp
+
 lemmas [code func del] = index.recs index.cases
 
 primrec
@@ -76,6 +80,8 @@
 lemma zero_index_code [code inline, code func]:
   "(0\<Colon>index) = Numeral0"
   by (simp add: number_of_index_def Pls_def)
+lemma [code post]: "Numeral0 = (0\<Colon>index)"
+  using zero_index_code ..
 
 definition [simp, code func del]:
   "(1\<Colon>index) = index_of_nat 1"
@@ -83,6 +89,8 @@
 lemma one_index_code [code inline, code func]:
   "(1\<Colon>index) = Numeral1"
   by (simp add: number_of_index_def Pls_def Bit_def)
+lemma [code post]: "Numeral1 = (1\<Colon>index)"
+  using one_index_code ..
 
 definition [simp, code func del]:
   "n + m = index_of_nat (nat_of_index n + nat_of_index m)"
@@ -179,7 +187,7 @@
 code_type index
   (SML "int")
   (OCaml "int")
-  (Haskell "Integer")
+  (Haskell "Int")
 
 code_instance index :: eq
   (Haskell -)
@@ -194,7 +202,7 @@
 
 code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index"
   (SML "Int.+/ ((_),/ (_))")
-  (OCaml "Pervasives.(+)")
+  (OCaml "Pervasives.( + )")
   (Haskell infixl 6 "+")
 
 code_const "op - \<Colon> index \<Rightarrow> index \<Rightarrow> index"
@@ -204,12 +212,12 @@
 
 code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index"
   (SML "Int.*/ ((_),/ (_))")
-  (OCaml "Pervasives.(*)")
+  (OCaml "Pervasives.( * )")
   (Haskell infixl 7 "*")
 
 code_const "op div \<Colon> index \<Rightarrow> index \<Rightarrow> index"
   (SML "Int.div/ ((_),/ (_))")
-  (OCaml "Pervasives.(/)")
+  (OCaml "Pervasives.( / )")
   (Haskell "div")
 
 code_const "op mod \<Colon> index \<Rightarrow> index \<Rightarrow> index"
@@ -219,17 +227,17 @@
 
 code_const "op = \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
   (SML "!((_ : Int.int) = _)")
-  (OCaml "!((_ : Pervasives.int) = _)")
+  (OCaml "!((_ : int) = _)")
   (Haskell infixl 4 "==")
 
 code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
   (SML "Int.<=/ ((_),/ (_))")
-  (OCaml "!((_ : Pervasives.int) <= _)")
+  (OCaml "!((_ : int) <= _)")
   (Haskell infix 4 "<=")
 
 code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
   (SML "Int.</ ((_),/ (_))")
-  (OCaml "!((_ : Pervasives.int) < _)")
+  (OCaml "!((_ : int) < _)")
   (Haskell infix 4 "<")
 
 end
--- a/src/HOL/Library/Efficient_Nat.thy	Fri Jan 25 14:54:44 2008 +0100
+++ b/src/HOL/Library/Efficient_Nat.thy	Fri Jan 25 14:54:46 2008 +0100
@@ -207,11 +207,14 @@
     then remove_suc_clause thy ths else ths
   end;
 
-fun lift_obj_eq f thy =
-  map (fn thm => thm RS @{thm meta_eq_to_obj_eq})
-  #> f thy
-  #> map (fn thm => thm RS @{thm eq_reflection})
-  #> map (Conv.fconv_rule Drule.beta_eta_conversion)
+fun lift_obj_eq f thy thms =
+  thms
+  |> try (
+    map (fn thm => thm RS @{thm meta_eq_to_obj_eq})
+    #> f thy
+    #> map (fn thm => thm RS @{thm eq_reflection})
+    #> map (Conv.fconv_rule Drule.beta_eta_conversion))
+  |> the_default thms
 *}
 
 setup {*
@@ -222,18 +225,16 @@
 *}
 (*>*)
 
-
 subsection {* Target language setup *}
 
 text {*
-  We map @{typ nat} to target language integers, where we
+  For ML, we map @{typ nat} to target language integers, where we
   assert that values are always non-negative.
 *}
 
 code_type nat
   (SML "int")
   (OCaml "Big'_int.big'_int")
-  (Haskell "Integer")
 
 types_code
   nat ("int")
@@ -247,22 +248,70 @@
 *}
 
 text {*
+  For Haskell we define our own @{typ nat} type.  The reason
+  is that we have to distinguish type class instances
+  for @{typ nat} and @{typ int}.
+*}
+
+code_include Haskell "Nat" {*
+newtype Nat = Nat Integer deriving (Show, Eq);
+
+instance Num Nat where {
+  fromInteger k = Nat (if k >= 0 then k else 0);
+  Nat n + Nat m = Nat (n + m);
+  Nat n - Nat m = fromInteger (n - m);
+  Nat n * Nat m = Nat (n * m);
+  abs n = n;
+  signum _ = 1;
+  negate n = error "negate Nat";
+};
+
+instance Ord Nat where {
+  Nat n <= Nat m = n <= m;
+  Nat n < Nat m = n < m;
+};
+
+instance Real Nat where {
+  toRational (Nat n) = toRational n;
+};
+
+instance Enum Nat where {
+  toEnum k = fromInteger (toEnum k);
+  fromEnum (Nat n) = fromEnum n;
+};
+
+instance Integral Nat where {
+  toInteger (Nat n) = n;
+  divMod n m = quotRem n m;
+  quotRem (Nat n) (Nat m) = (Nat k, Nat l) where (k, l) = quotRem n m;
+};
+*}
+
+code_reserved Haskell Nat
+
+code_type nat
+  (Haskell "Nat")
+
+code_instance nat :: eq
+  (Haskell -)
+
+text {*
   Natural numerals.
 *}
 
-lemma [code inline]:
+lemma [code inline, symmetric, code post]:
   "nat (number_of i) = number_nat_inst.number_of_nat i"
   -- {* this interacts as desired with @{thm nat_number_of_def} *}
   by (simp add: number_nat_inst.number_of_nat)
 
 setup {*
   fold (Numeral.add_code @{const_name number_nat_inst.number_of_nat}
-    false false) ["SML", "OCaml", "Haskell"]
+    true false) ["SML", "OCaml", "Haskell"]
 *}
 
 text {*
   Since natural numbers are implemented
-  using integers, the coercion function @{const "of_nat"} of type
+  using integers in ML, the coercion function @{const "of_nat"} of type
   @{typ "nat \<Rightarrow> int"} is simply implemented by the identity function.
   For the @{const "nat"} function for converting an integer to a natural
   number, we give a specific implementation using an ML function that
@@ -285,16 +334,11 @@
 
 lemma of_nat_int [code unfold]:
   "of_nat = int" by (simp add: int_def)
+declare of_nat_int [symmetric, code post]
 
 code_const int
   (SML "_")
   (OCaml "_")
-  (Haskell "_")
-
-code_const nat
-  (SML "IntInf.max/ (/0,/ _)")
-  (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int")
-  (Haskell "max 0")
 
 consts_code
   int ("(_)")
@@ -303,19 +347,26 @@
 fun nat i = if i < 0 then 0 else i;
 *}
 
+code_const nat
+  (SML "IntInf.max/ (/0,/ _)")
+  (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int")
+
+text {* For Haskell, things are slightly different again. *}
+
+code_const int and nat
+  (Haskell "toInteger" and "fromInteger")
 
 text {* Conversion from and to indices. *}
 
+code_const index_of_nat
+  (SML "IntInf.toInt")
+  (OCaml "Big'_int.int'_of'_big'_int")
+  (Haskell "toEnum")
+
 code_const nat_of_index
   (SML "IntInf.fromInt")
   (OCaml "Big'_int.big'_int'_of'_int")
-  (Haskell "toInteger")
-
-code_const index_of_nat
-  (SML "IntInf.toInt")
-  (OCaml "Big'_int.int'_of'_big'_int")
-  (Haskell "fromInteger")
-
+  (Haskell "fromEnum")
 
 text {* Using target language arithmetic operations whenever appropriate *}
 
--- a/src/HOL/Tools/numeral.ML	Fri Jan 25 14:54:44 2008 +0100
+++ b/src/HOL/Tools/numeral.ML	Fri Jan 25 14:54:46 2008 +0100
@@ -56,7 +56,7 @@
 (* code generator *)
 
 fun add_code number_of negative unbounded target =
-  CodeTarget.add_pretty_numeral target unbounded negative number_of
+  CodeTarget.add_pretty_numeral target negative unbounded number_of
   @{const_name Int.B0} @{const_name Int.B1}
   @{const_name Int.Pls} @{const_name Int.Min}
   @{const_name Int.Bit};