| author | blanchet |
| Tue, 31 Aug 2010 20:24:28 +0200 | |
| changeset 38942 | e10c11971fa7 |
| parent 38857 | 97775f3e8722 |
| child 38968 | e55deaa22fff |
| permissions | -rw-r--r-- |
| 37968 | 1 |
(* Title: HOL/Library/Code_Natural.thy |
| 24999 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
*) |
|
4 |
||
| 37968 | 5 |
theory Code_Natural |
|
31203
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
haftmann
parents:
31192
diff
changeset
|
6 |
imports Main |
| 24999 | 7 |
begin |
8 |
||
| 37968 | 9 |
section {* Alternative representation of @{typ code_numeral} for @{text Haskell} and @{text Scala} *}
|
| 24999 | 10 |
|
| 37968 | 11 |
code_include Haskell "Natural" |
| 38810 | 12 |
{*import Data.Array.ST;
|
13 |
||
14 |
newtype Natural = Natural Integer deriving (Eq, Show, Read); |
|
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
15 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
16 |
instance Num Natural where {
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
17 |
fromInteger k = Natural (if k >= 0 then k else 0); |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
18 |
Natural n + Natural m = Natural (n + m); |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
19 |
Natural n - Natural m = fromInteger (n - m); |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
20 |
Natural n * Natural m = Natural (n * m); |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
21 |
abs n = n; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
22 |
signum _ = 1; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
23 |
negate n = error "negate Natural"; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
24 |
}; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
25 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
26 |
instance Ord Natural where {
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
27 |
Natural n <= Natural m = n <= m; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
28 |
Natural n < Natural m = n < m; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
29 |
}; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
30 |
|
| 37968 | 31 |
instance Ix Natural where {
|
32 |
range (Natural n, Natural m) = map Natural (range (n, m)); |
|
33 |
index (Natural n, Natural m) (Natural q) = index (n, m) q; |
|
34 |
inRange (Natural n, Natural m) (Natural q) = inRange (n, m) q; |
|
35 |
rangeSize (Natural n, Natural m) = rangeSize (n, m); |
|
36 |
}; |
|
37 |
||
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
38 |
instance Real Natural where {
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
39 |
toRational (Natural n) = toRational n; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
40 |
}; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
41 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
42 |
instance Enum Natural where {
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
43 |
toEnum k = fromInteger (toEnum k); |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
44 |
fromEnum (Natural n) = fromEnum n; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
45 |
}; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
46 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
47 |
instance Integral Natural where {
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
48 |
toInteger (Natural n) = n; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
49 |
divMod n m = quotRem n m; |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
50 |
quotRem (Natural n) (Natural m) |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
51 |
| (m == 0) = (0, Natural n) |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
52 |
| otherwise = (Natural k, Natural l) where (k, l) = quotRem n m; |
| 37968 | 53 |
};*} |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
54 |
|
| 38810 | 55 |
|
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
56 |
code_reserved Haskell Natural |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
57 |
|
| 38774 | 58 |
code_include Scala "Natural" |
59 |
{*import scala.Math
|
|
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
60 |
|
| 38771 | 61 |
object Nat {
|
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
62 |
|
| 38771 | 63 |
def apply(numeral: BigInt): Nat = new Nat(numeral max 0) |
64 |
def apply(numeral: Int): Nat = Nat(BigInt(numeral)) |
|
65 |
def apply(numeral: String): Nat = Nat(BigInt(numeral)) |
|
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
66 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
67 |
} |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
68 |
|
| 38771 | 69 |
class Nat private(private val value: BigInt) {
|
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
70 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
71 |
override def hashCode(): Int = this.value.hashCode() |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
72 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
73 |
override def equals(that: Any): Boolean = that match {
|
| 38771 | 74 |
case that: Nat => this equals that |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
75 |
case _ => false |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
76 |
} |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
77 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
78 |
override def toString(): String = this.value.toString |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
79 |
|
| 38771 | 80 |
def equals(that: Nat): Boolean = this.value == that.value |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
81 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
82 |
def as_BigInt: BigInt = this.value |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
83 |
def as_Int: Int = if (this.value >= Int.MinValue && this.value <= Int.MaxValue) |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
84 |
this.value.intValue |
| 37968 | 85 |
else error("Int value out of range: " + this.value.toString)
|
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
86 |
|
| 38771 | 87 |
def +(that: Nat): Nat = new Nat(this.value + that.value) |
88 |
def -(that: Nat): Nat = Nat(this.value - that.value) |
|
89 |
def *(that: Nat): Nat = new Nat(this.value * that.value) |
|
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
90 |
|
| 38771 | 91 |
def /%(that: Nat): (Nat, Nat) = if (that.value == 0) (new Nat(0), this) |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
92 |
else {
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
93 |
val (k, l) = this.value /% that.value |
| 38771 | 94 |
(new Nat(k), new Nat(l)) |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
95 |
} |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
96 |
|
| 38771 | 97 |
def <=(that: Nat): Boolean = this.value <= that.value |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
98 |
|
| 38771 | 99 |
def <(that: Nat): Boolean = this.value < that.value |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
100 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
101 |
} |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
102 |
*} |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
103 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
104 |
code_reserved Scala Natural |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
105 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
106 |
code_type code_numeral |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
107 |
(Haskell "Natural.Natural") |
| 38771 | 108 |
(Scala "Natural.Nat") |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
109 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
110 |
setup {*
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
111 |
fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
112 |
false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"] |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
113 |
*} |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
114 |
|
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38810
diff
changeset
|
115 |
code_instance code_numeral :: equal |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
116 |
(Haskell -) |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
117 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
118 |
code_const "op + \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
119 |
(Haskell infixl 6 "+") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
120 |
(Scala infixl 7 "+") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
121 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
122 |
code_const "op - \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
123 |
(Haskell infixl 6 "-") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
124 |
(Scala infixl 7 "-") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
125 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
126 |
code_const "op * \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
127 |
(Haskell infixl 7 "*") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
128 |
(Scala infixl 8 "*") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
129 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
130 |
code_const div_mod_code_numeral |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
131 |
(Haskell "divMod") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
132 |
(Scala infixl 8 "/%") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
133 |
|
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38810
diff
changeset
|
134 |
code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
135 |
(Haskell infixl 4 "==") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
136 |
(Scala infixl 5 "==") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
137 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
138 |
code_const "op \<le> \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
139 |
(Haskell infix 4 "<=") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
140 |
(Scala infixl 4 "<=") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
141 |
|
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
142 |
code_const "op < \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
143 |
(Haskell infix 4 "<") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
144 |
(Scala infixl 4 "<") |
|
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
145 |
|
| 37968 | 146 |
end |