1 (* Title: HOL/Library/Code_Natural.thy
2 Author: Florian Haftmann, TU Muenchen
9 section {* Alternative representation of @{typ code_numeral} for @{text Haskell} and @{text Scala} *}
11 code_include Haskell "Natural"
12 {*newtype Natural = Natural Integer deriving (Eq, Show, Read);
14 instance Num Natural where {
15 fromInteger k = Natural (if k >= 0 then k else 0);
16 Natural n + Natural m = Natural (n + m);
17 Natural n - Natural m = fromInteger (n - m);
18 Natural n * Natural m = Natural (n * m);
21 negate n = error "negate Natural";
24 instance Ord Natural where {
25 Natural n <= Natural m = n <= m;
26 Natural n < Natural m = n < m;
29 instance Ix Natural where {
30 range (Natural n, Natural m) = map Natural (range (n, m));
31 index (Natural n, Natural m) (Natural q) = index (n, m) q;
32 inRange (Natural n, Natural m) (Natural q) = inRange (n, m) q;
33 rangeSize (Natural n, Natural m) = rangeSize (n, m);
36 instance Real Natural where {
37 toRational (Natural n) = toRational n;
40 instance Enum Natural where {
41 toEnum k = fromInteger (toEnum k);
42 fromEnum (Natural n) = fromEnum n;
45 instance Integral Natural where {
46 toInteger (Natural n) = n;
47 divMod n m = quotRem n m;
48 quotRem (Natural n) (Natural m)
49 | (m == 0) = (0, Natural n)
50 | otherwise = (Natural k, Natural l) where (k, l) = quotRem n m;
53 code_reserved Haskell Natural
55 code_include Scala "Natural"
60 def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
61 def apply(numeral: Int): Nat = Nat(BigInt(numeral))
62 def apply(numeral: String): Nat = Nat(BigInt(numeral))
66 class Nat private(private val value: BigInt) {
68 override def hashCode(): Int = this.value.hashCode()
70 override def equals(that: Any): Boolean = that match {
71 case that: Nat => this equals that
75 override def toString(): String = this.value.toString
77 def equals(that: Nat): Boolean = this.value == that.value
79 def as_BigInt: BigInt = this.value
80 def as_Int: Int = if (this.value >= Int.MinValue && this.value <= Int.MaxValue)
82 else error("Int value out of range: " + this.value.toString)
84 def +(that: Nat): Nat = new Nat(this.value + that.value)
85 def -(that: Nat): Nat = Nat(this.value - that.value)
86 def *(that: Nat): Nat = new Nat(this.value * that.value)
88 def /%(that: Nat): (Nat, Nat) = if (that.value == 0) (new Nat(0), this)
90 val (k, l) = this.value /% that.value
91 (new Nat(k), new Nat(l))
94 def <=(that: Nat): Boolean = this.value <= that.value
96 def <(that: Nat): Boolean = this.value < that.value
101 code_reserved Scala Natural
103 code_type code_numeral
104 (Haskell "Natural.Natural")
105 (Scala "Natural.Nat")
108 fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
109 false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"]
112 code_instance code_numeral :: eq
115 code_const "op + \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
116 (Haskell infixl 6 "+")
119 code_const "op - \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
120 (Haskell infixl 6 "-")
123 code_const "op * \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
124 (Haskell infixl 7 "*")
127 code_const div_mod_code_numeral
129 (Scala infixl 8 "/%")
131 code_const "eq_class.eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
132 (Haskell infixl 4 "==")
133 (Scala infixl 5 "==")
135 code_const "op \<le> \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
136 (Haskell infix 4 "<=")
137 (Scala infixl 4 "<=")
139 code_const "op < \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
140 (Haskell infix 4 "<")