--- a/src/HOL/Library/Code_Natural.thy Wed Sep 01 09:03:34 2010 +0200
+++ b/src/HOL/Library/Code_Natural.thy Wed Sep 01 11:09:50 2010 +0200
@@ -56,47 +56,45 @@
code_reserved Haskell Natural
code_include Scala "Natural"
-{*import scala.Math
-
-object Nat {
+{*object Natural {
- def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
- def apply(numeral: Int): Nat = Nat(BigInt(numeral))
- def apply(numeral: String): Nat = Nat(BigInt(numeral))
+ def apply(numeral: BigInt): Natural = new Natural(numeral max 0)
+ def apply(numeral: Int): Natural = Natural(BigInt(numeral))
+ def apply(numeral: String): Natural = Natural(BigInt(numeral))
}
-class Nat private(private val value: BigInt) {
+class Natural private(private val value: BigInt) {
override def hashCode(): Int = this.value.hashCode()
override def equals(that: Any): Boolean = that match {
- case that: Nat => this equals that
+ case that: Natural => this equals that
case _ => false
}
override def toString(): String = this.value.toString
- def equals(that: Nat): Boolean = this.value == that.value
+ def equals(that: Natural): Boolean = this.value == that.value
def as_BigInt: BigInt = this.value
def as_Int: Int = if (this.value >= Int.MinValue && this.value <= Int.MaxValue)
this.value.intValue
else error("Int value out of range: " + this.value.toString)
- def +(that: Nat): Nat = new Nat(this.value + that.value)
- def -(that: Nat): Nat = Nat(this.value - that.value)
- def *(that: Nat): Nat = new Nat(this.value * that.value)
+ def +(that: Natural): Natural = new Natural(this.value + that.value)
+ def -(that: Natural): Natural = Natural(this.value - that.value)
+ def *(that: Natural): Natural = new Natural(this.value * that.value)
- def /%(that: Nat): (Nat, Nat) = if (that.value == 0) (new Nat(0), this)
+ def /%(that: Natural): (Natural, Natural) = if (that.value == 0) (new Natural(0), this)
else {
val (k, l) = this.value /% that.value
- (new Nat(k), new Nat(l))
+ (new Natural(k), new Natural(l))
}
- def <=(that: Nat): Boolean = this.value <= that.value
+ def <=(that: Natural): Boolean = this.value <= that.value
- def <(that: Nat): Boolean = this.value < that.value
+ def <(that: Natural): Boolean = this.value < that.value
}
*}
@@ -105,7 +103,7 @@
code_type code_numeral
(Haskell "Natural.Natural")
- (Scala "Natural.Nat")
+ (Scala "Natural")
setup {*
fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}