haftmann@37968: (* Title: HOL/Library/Code_Natural.thy haftmann@24999: Author: Florian Haftmann, TU Muenchen haftmann@24999: *) haftmann@24999: haftmann@37968: theory Code_Natural haftmann@31203: imports Main haftmann@24999: begin haftmann@24999: haftmann@37968: section {* Alternative representation of @{typ code_numeral} for @{text Haskell} and @{text Scala} *} haftmann@24999: haftmann@37968: code_include Haskell "Natural" haftmann@38810: {*import Data.Array.ST; haftmann@38810: haftmann@38810: newtype Natural = Natural Integer deriving (Eq, Show, Read); haftmann@37958: haftmann@37958: instance Num Natural where { haftmann@37958: fromInteger k = Natural (if k >= 0 then k else 0); haftmann@37958: Natural n + Natural m = Natural (n + m); haftmann@37958: Natural n - Natural m = fromInteger (n - m); haftmann@37958: Natural n * Natural m = Natural (n * m); haftmann@37958: abs n = n; haftmann@37958: signum _ = 1; haftmann@37958: negate n = error "negate Natural"; haftmann@37958: }; haftmann@37958: haftmann@37958: instance Ord Natural where { haftmann@37958: Natural n <= Natural m = n <= m; haftmann@37958: Natural n < Natural m = n < m; haftmann@37958: }; haftmann@37958: haftmann@37968: instance Ix Natural where { haftmann@37968: range (Natural n, Natural m) = map Natural (range (n, m)); haftmann@37968: index (Natural n, Natural m) (Natural q) = index (n, m) q; haftmann@37968: inRange (Natural n, Natural m) (Natural q) = inRange (n, m) q; haftmann@37968: rangeSize (Natural n, Natural m) = rangeSize (n, m); haftmann@37968: }; haftmann@37968: haftmann@37958: instance Real Natural where { haftmann@37958: toRational (Natural n) = toRational n; haftmann@37958: }; haftmann@37958: haftmann@37958: instance Enum Natural where { haftmann@37958: toEnum k = fromInteger (toEnum k); haftmann@37958: fromEnum (Natural n) = fromEnum n; haftmann@37958: }; haftmann@37958: haftmann@37958: instance Integral Natural where { haftmann@37958: toInteger (Natural n) = n; haftmann@37958: divMod n m = quotRem n m; haftmann@37958: quotRem (Natural n) (Natural m) haftmann@37958: | (m == 0) = (0, Natural n) haftmann@37958: | otherwise = (Natural k, Natural l) where (k, l) = quotRem n m; haftmann@37968: };*} haftmann@37958: haftmann@38810: haftmann@37958: code_reserved Haskell Natural haftmann@37958: haftmann@38774: code_include Scala "Natural" haftmann@38968: {*object Natural { haftmann@37958: haftmann@38968: def apply(numeral: BigInt): Natural = new Natural(numeral max 0) haftmann@38968: def apply(numeral: Int): Natural = Natural(BigInt(numeral)) haftmann@38968: def apply(numeral: String): Natural = Natural(BigInt(numeral)) haftmann@37958: haftmann@37958: } haftmann@37958: haftmann@38968: class Natural private(private val value: BigInt) { haftmann@37958: haftmann@37958: override def hashCode(): Int = this.value.hashCode() haftmann@37958: haftmann@37958: override def equals(that: Any): Boolean = that match { haftmann@38968: case that: Natural => this equals that haftmann@37958: case _ => false haftmann@37958: } haftmann@37958: haftmann@37958: override def toString(): String = this.value.toString haftmann@37958: haftmann@38968: def equals(that: Natural): Boolean = this.value == that.value haftmann@37958: haftmann@37958: def as_BigInt: BigInt = this.value haftmann@37958: def as_Int: Int = if (this.value >= Int.MinValue && this.value <= Int.MaxValue) haftmann@37958: this.value.intValue haftmann@37968: else error("Int value out of range: " + this.value.toString) haftmann@37958: haftmann@38968: def +(that: Natural): Natural = new Natural(this.value + that.value) haftmann@38968: def -(that: Natural): Natural = Natural(this.value - that.value) haftmann@38968: def *(that: Natural): Natural = new Natural(this.value * that.value) haftmann@37958: haftmann@38968: def /%(that: Natural): (Natural, Natural) = if (that.value == 0) (new Natural(0), this) haftmann@37958: else { haftmann@37958: val (k, l) = this.value /% that.value haftmann@38968: (new Natural(k), new Natural(l)) haftmann@37958: } haftmann@37958: haftmann@38968: def <=(that: Natural): Boolean = this.value <= that.value haftmann@37958: haftmann@38968: def <(that: Natural): Boolean = this.value < that.value haftmann@37958: haftmann@37958: } haftmann@37958: *} haftmann@37958: haftmann@37958: code_reserved Scala Natural haftmann@37958: haftmann@37958: code_type code_numeral haftmann@37958: (Haskell "Natural.Natural") haftmann@38968: (Scala "Natural") haftmann@37958: haftmann@37958: setup {* haftmann@37958: fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} haftmann@37958: false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"] haftmann@37958: *} haftmann@37958: haftmann@38857: code_instance code_numeral :: equal haftmann@37958: (Haskell -) haftmann@37958: haftmann@37958: code_const "op + \ code_numeral \ code_numeral \ code_numeral" haftmann@37958: (Haskell infixl 6 "+") haftmann@37958: (Scala infixl 7 "+") haftmann@37958: haftmann@37958: code_const "op - \ code_numeral \ code_numeral \ code_numeral" haftmann@37958: (Haskell infixl 6 "-") haftmann@37958: (Scala infixl 7 "-") haftmann@37958: haftmann@37958: code_const "op * \ code_numeral \ code_numeral \ code_numeral" haftmann@37958: (Haskell infixl 7 "*") haftmann@37958: (Scala infixl 8 "*") haftmann@37958: haftmann@37958: code_const div_mod_code_numeral haftmann@37958: (Haskell "divMod") haftmann@37958: (Scala infixl 8 "/%") haftmann@37958: haftmann@38857: code_const "HOL.equal \ code_numeral \ code_numeral \ bool" haftmann@37958: (Haskell infixl 4 "==") haftmann@37958: (Scala infixl 5 "==") haftmann@37958: haftmann@37958: code_const "op \ \ code_numeral \ code_numeral \ bool" haftmann@37958: (Haskell infix 4 "<=") haftmann@37958: (Scala infixl 4 "<=") haftmann@37958: haftmann@37958: code_const "op < \ code_numeral \ code_numeral \ bool" haftmann@37958: (Haskell infix 4 "<") haftmann@37958: (Scala infixl 4 "<") haftmann@37958: haftmann@37968: end