src/HOL/Library/Efficient_Nat.thy
author haftmann
Mon, 23 Jul 2012 09:28:03 +0200
changeset 49446 6efff142bb54
parent 49088 1b609a7837ef
child 49583 084cd758a8ab
permissions -rw-r--r--
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
     1 (*  Title:      HOL/Library/Efficient_Nat.thy
     2     Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
     3 *)
     4 
     5 header {* Implementation of natural numbers by target-language integers *}
     6 
     7 theory Efficient_Nat
     8 imports Code_Nat Code_Integer Main
     9 begin
    10 
    11 text {*
    12   The efficiency of the generated code for natural numbers can be improved
    13   drastically by implementing natural numbers by target-language
    14   integers.  To do this, just include this theory.
    15 *}
    16 
    17 subsection {* Target language fundamentals *}
    18 
    19 text {*
    20   For ML, we map @{typ nat} to target language integers, where we
    21   ensure that values are always non-negative.
    22 *}
    23 
    24 code_type nat
    25   (SML "IntInf.int")
    26   (OCaml "Big'_int.big'_int")
    27   (Eval "int")
    28 
    29 text {*
    30   For Haskell and Scala we define our own @{typ nat} type.  The reason
    31   is that we have to distinguish type class instances for @{typ nat}
    32   and @{typ int}.
    33 *}
    34 
    35 code_include Haskell "Nat"
    36 {*newtype Nat = Nat Integer deriving (Eq, Show, Read);
    37 
    38 instance Num Nat where {
    39   fromInteger k = Nat (if k >= 0 then k else 0);
    40   Nat n + Nat m = Nat (n + m);
    41   Nat n - Nat m = fromInteger (n - m);
    42   Nat n * Nat m = Nat (n * m);
    43   abs n = n;
    44   signum _ = 1;
    45   negate n = error "negate Nat";
    46 };
    47 
    48 instance Ord Nat where {
    49   Nat n <= Nat m = n <= m;
    50   Nat n < Nat m = n < m;
    51 };
    52 
    53 instance Real Nat where {
    54   toRational (Nat n) = toRational n;
    55 };
    56 
    57 instance Enum Nat where {
    58   toEnum k = fromInteger (toEnum k);
    59   fromEnum (Nat n) = fromEnum n;
    60 };
    61 
    62 instance Integral Nat where {
    63   toInteger (Nat n) = n;
    64   divMod n m = quotRem n m;
    65   quotRem (Nat n) (Nat m)
    66     | (m == 0) = (0, Nat n)
    67     | otherwise = (Nat k, Nat l) where (k, l) = quotRem n m;
    68 };
    69 *}
    70 
    71 code_reserved Haskell Nat
    72 
    73 code_include Scala "Nat"
    74 {*object Nat {
    75 
    76   def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
    77   def apply(numeral: Int): Nat = Nat(BigInt(numeral))
    78   def apply(numeral: String): Nat = Nat(BigInt(numeral))
    79 
    80 }
    81 
    82 class Nat private(private val value: BigInt) {
    83 
    84   override def hashCode(): Int = this.value.hashCode()
    85 
    86   override def equals(that: Any): Boolean = that match {
    87     case that: Nat => this equals that
    88     case _ => false
    89   }
    90 
    91   override def toString(): String = this.value.toString
    92 
    93   def equals(that: Nat): Boolean = this.value == that.value
    94 
    95   def as_BigInt: BigInt = this.value
    96   def as_Int: Int = if (this.value >= scala.Int.MinValue && this.value <= scala.Int.MaxValue)
    97       this.value.intValue
    98     else error("Int value out of range: " + this.value.toString)
    99 
   100   def +(that: Nat): Nat = new Nat(this.value + that.value)
   101   def -(that: Nat): Nat = Nat(this.value - that.value)
   102   def *(that: Nat): Nat = new Nat(this.value * that.value)
   103 
   104   def /%(that: Nat): (Nat, Nat) = if (that.value == 0) (new Nat(0), this)
   105     else {
   106       val (k, l) = this.value /% that.value
   107       (new Nat(k), new Nat(l))
   108     }
   109 
   110   def <=(that: Nat): Boolean = this.value <= that.value
   111 
   112   def <(that: Nat): Boolean = this.value < that.value
   113 
   114 }
   115 *}
   116 
   117 code_reserved Scala Nat
   118 
   119 code_type nat
   120   (Haskell "Nat.Nat")
   121   (Scala "Nat")
   122 
   123 code_instance nat :: equal
   124   (Haskell -)
   125 
   126 setup {*
   127   fold (Numeral.add_code @{const_name nat_of_num}
   128     false Code_Printer.literal_positive_numeral) ["SML", "OCaml", "Haskell", "Scala"]
   129 *}
   130 
   131 code_const "0::nat"
   132   (SML "0")
   133   (OCaml "Big'_int.zero'_big'_int")
   134   (Haskell "0")
   135   (Scala "Nat(0)")
   136 
   137 
   138 subsection {* Conversions *}
   139 
   140 text {*
   141   Since natural numbers are implemented
   142   using integers in ML, the coercion function @{term "int"} of type
   143   @{typ "nat \<Rightarrow> int"} is simply implemented by the identity function.
   144   For the @{const nat} function for converting an integer to a natural
   145   number, we give a specific implementation using an ML expression that
   146   returns its input value, provided that it is non-negative, and otherwise
   147   returns @{text "0"}.
   148 *}
   149 
   150 definition int :: "nat \<Rightarrow> int" where
   151   [code_abbrev]: "int = of_nat"
   152 
   153 code_const int
   154   (SML "_")
   155   (OCaml "_")
   156 
   157 code_const nat
   158   (SML "IntInf.max/ (0,/ _)")
   159   (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int")
   160   (Eval "Integer.max/ 0")
   161 
   162 text {* For Haskell and Scala, things are slightly different again. *}
   163 
   164 code_const int and nat
   165   (Haskell "Prelude.toInteger" and "Prelude.fromInteger")
   166   (Scala "!_.as'_BigInt" and "Nat")
   167 
   168 text {* Alternativ implementation for @{const of_nat} *}
   169 
   170 lemma [code]:
   171   "of_nat n = (if n = 0 then 0 else
   172      let
   173        (q, m) = divmod_nat n 2;
   174        q' = 2 * of_nat q
   175      in if m = 0 then q' else q' + 1)"
   176 proof -
   177   from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp
   178   show ?thesis
   179     apply (simp add: Let_def divmod_nat_div_mod mod_2_not_eq_zero_eq_one_nat
   180       of_nat_mult
   181       of_nat_add [symmetric])
   182     apply (auto simp add: of_nat_mult)
   183     apply (simp add: * of_nat_mult add_commute mult_commute)
   184     done
   185 qed
   186 
   187 text {* Conversion from and to code numerals *}
   188 
   189 code_const Code_Numeral.of_nat
   190   (SML "IntInf.toInt")
   191   (OCaml "_")
   192   (Haskell "!(Prelude.fromInteger/ ./ Prelude.toInteger)")
   193   (Scala "!Natural(_.as'_BigInt)")
   194   (Eval "_")
   195 
   196 code_const Code_Numeral.nat_of
   197   (SML "IntInf.fromInt")
   198   (OCaml "_")
   199   (Haskell "!(Prelude.fromInteger/ ./ Prelude.toInteger)")
   200   (Scala "!Nat(_.as'_BigInt)")
   201   (Eval "_")
   202 
   203 
   204 subsection {* Target language arithmetic *}
   205 
   206 code_const "plus \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
   207   (SML "IntInf.+/ ((_),/ (_))")
   208   (OCaml "Big'_int.add'_big'_int")
   209   (Haskell infixl 6 "+")
   210   (Scala infixl 7 "+")
   211   (Eval infixl 8 "+")
   212 
   213 code_const "minus \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
   214   (SML "IntInf.max/ (0, IntInf.-/ ((_),/ (_)))")
   215   (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)")
   216   (Haskell infixl 6 "-")
   217   (Scala infixl 7 "-")
   218   (Eval "Integer.max/ 0/ (_ -/ _)")
   219 
   220 code_const Code_Nat.dup
   221   (SML "IntInf.*/ (2,/ (_))")
   222   (OCaml "Big'_int.mult'_big'_int/ 2")
   223   (Haskell "!(2 * _)")
   224   (Scala "!(2 * _)")
   225   (Eval "!(2 * _)")
   226 
   227 code_const Code_Nat.sub
   228   (SML "!(raise/ Fail/ \"sub\")")
   229   (OCaml "failwith/ \"sub\"")
   230   (Haskell "error/ \"sub\"")
   231   (Scala "!sys.error(\"sub\")")
   232 
   233 code_const "times \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
   234   (SML "IntInf.*/ ((_),/ (_))")
   235   (OCaml "Big'_int.mult'_big'_int")
   236   (Haskell infixl 7 "*")
   237   (Scala infixl 8 "*")
   238   (Eval infixl 9 "*")
   239 
   240 code_const divmod_nat
   241   (SML "IntInf.divMod/ ((_),/ (_))")
   242   (OCaml "Big'_int.quomod'_big'_int")
   243   (Haskell "divMod")
   244   (Scala infixl 8 "/%")
   245   (Eval "Integer.div'_mod")
   246 
   247 code_const "HOL.equal \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
   248   (SML "!((_ : IntInf.int) = _)")
   249   (OCaml "Big'_int.eq'_big'_int")
   250   (Haskell infix 4 "==")
   251   (Scala infixl 5 "==")
   252   (Eval infixl 6 "=")
   253 
   254 code_const "less_eq \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
   255   (SML "IntInf.<=/ ((_),/ (_))")
   256   (OCaml "Big'_int.le'_big'_int")
   257   (Haskell infix 4 "<=")
   258   (Scala infixl 4 "<=")
   259   (Eval infixl 6 "<=")
   260 
   261 code_const "less \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
   262   (SML "IntInf.</ ((_),/ (_))")
   263   (OCaml "Big'_int.lt'_big'_int")
   264   (Haskell infix 4 "<")
   265   (Scala infixl 4 "<")
   266   (Eval infixl 6 "<")
   267 
   268 code_const Num.num_of_nat
   269   (SML "!(raise/ Fail/ \"num'_of'_nat\")")
   270   (OCaml "failwith/ \"num'_of'_nat\"")
   271   (Haskell "error/ \"num'_of'_nat\"")
   272   (Scala "!sys.error(\"num'_of'_nat\")")
   273 
   274 
   275 subsection {* Evaluation *}
   276 
   277 lemma [code, code del]:
   278   "(Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term) = Code_Evaluation.term_of" ..
   279 
   280 code_const "Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term"
   281   (SML "HOLogic.mk'_number/ HOLogic.natT")
   282 
   283 text {*
   284   FIXME -- Evaluation with @{text "Quickcheck_Narrowing"} does not work, as
   285   @{text "code_module"} is very aggressive leading to bad Haskell code.
   286   Therefore, we simply deactivate the narrowing-based quickcheck from here on.
   287 *}
   288 
   289 declare [[quickcheck_narrowing_active = false]] 
   290 
   291 
   292 code_modulename SML
   293   Efficient_Nat Arith
   294 
   295 code_modulename OCaml
   296   Efficient_Nat Arith
   297 
   298 code_modulename Haskell
   299   Efficient_Nat Arith
   300 
   301 hide_const (open) int
   302 
   303 end