| author | wenzelm | 
| Sun, 12 Jun 2011 16:19:29 +0200 | |
| changeset 43368 | 0dc67b3cf8a5 | 
| parent 39781 | 2053638a2bf2 | 
| child 46547 | d1dcb91a512e | 
| 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: 
31192diff
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: 
37947diff
changeset | 15 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 16 | instance Num Natural where {
 | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 17 | fromInteger k = Natural (if k >= 0 then k else 0); | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 18 | Natural n + Natural m = Natural (n + m); | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 19 | Natural n - Natural m = fromInteger (n - m); | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 20 | Natural n * Natural m = Natural (n * m); | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 21 | abs n = n; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 22 | signum _ = 1; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 23 | negate n = error "negate Natural"; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 24 | }; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 25 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 26 | instance Ord Natural where {
 | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 27 | Natural n <= Natural m = n <= m; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 28 | Natural n < Natural m = n < m; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 29 | }; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
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: 
37947diff
changeset | 38 | instance Real Natural where {
 | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 39 | toRational (Natural n) = toRational n; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 40 | }; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 41 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 42 | instance Enum Natural where {
 | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 43 | toEnum k = fromInteger (toEnum k); | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 44 | fromEnum (Natural n) = fromEnum n; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 45 | }; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 46 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 47 | instance Integral Natural where {
 | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 48 | toInteger (Natural n) = n; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 49 | divMod n m = quotRem n m; | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 50 | quotRem (Natural n) (Natural m) | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 51 | | (m == 0) = (0, Natural n) | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
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: 
37947diff
changeset | 54 | |
| 38810 | 55 | |
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 56 | code_reserved Haskell Natural | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 57 | |
| 38774 | 58 | code_include Scala "Natural" | 
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 59 | {*object Natural {
 | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 60 | |
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 61 | def apply(numeral: BigInt): Natural = new Natural(numeral max 0) | 
| 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 62 | def apply(numeral: Int): Natural = Natural(BigInt(numeral)) | 
| 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 63 | def apply(numeral: String): Natural = Natural(BigInt(numeral)) | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 64 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 65 | } | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 66 | |
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 67 | class Natural private(private val value: BigInt) {
 | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 68 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 69 | override def hashCode(): Int = this.value.hashCode() | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 70 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 71 |   override def equals(that: Any): Boolean = that match {
 | 
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 72 | case that: Natural => this equals that | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 73 | case _ => false | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 74 | } | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 75 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 76 | override def toString(): String = this.value.toString | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 77 | |
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 78 | def equals(that: Natural): Boolean = this.value == that.value | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 79 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 80 | def as_BigInt: BigInt = this.value | 
| 39781 | 81 | def as_Int: Int = if (this.value >= scala.Int.MinValue && this.value <= scala.Int.MaxValue) | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 82 | this.value.intValue | 
| 37968 | 83 |     else error("Int value out of range: " + this.value.toString)
 | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 84 | |
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 85 | def +(that: Natural): Natural = new Natural(this.value + that.value) | 
| 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 86 | def -(that: Natural): Natural = Natural(this.value - that.value) | 
| 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 87 | def *(that: Natural): Natural = new Natural(this.value * that.value) | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 88 | |
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 89 | def /%(that: Natural): (Natural, Natural) = if (that.value == 0) (new Natural(0), this) | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 90 |     else {
 | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 91 | val (k, l) = this.value /% that.value | 
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 92 | (new Natural(k), new Natural(l)) | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 93 | } | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 94 | |
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 95 | def <=(that: Natural): Boolean = this.value <= that.value | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 96 | |
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 97 | def <(that: Natural): Boolean = this.value < that.value | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 98 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 99 | } | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 100 | *} | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 101 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 102 | code_reserved Scala Natural | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 103 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 104 | code_type code_numeral | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 105 | (Haskell "Natural.Natural") | 
| 38968 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 haftmann parents: 
38857diff
changeset | 106 | (Scala "Natural") | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 107 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 108 | setup {*
 | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 109 |   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: 
37947diff
changeset | 110 | false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"] | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 111 | *} | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 112 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38810diff
changeset | 113 | code_instance code_numeral :: equal | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 114 | (Haskell -) | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 115 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 116 | code_const "op + \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 117 | (Haskell infixl 6 "+") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 118 | (Scala infixl 7 "+") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 119 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 120 | code_const "op - \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 121 | (Haskell infixl 6 "-") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 122 | (Scala infixl 7 "-") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 123 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 124 | code_const "op * \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 125 | (Haskell infixl 7 "*") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 126 | (Scala infixl 8 "*") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 127 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 128 | code_const div_mod_code_numeral | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 129 | (Haskell "divMod") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 130 | (Scala infixl 8 "/%") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 131 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38810diff
changeset | 132 | code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" | 
| 39272 | 133 | (Haskell infix 4 "==") | 
| 37958 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 134 | (Scala infixl 5 "==") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 135 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 136 | code_const "op \<le> \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 137 | (Haskell infix 4 "<=") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 138 | (Scala infixl 4 "<=") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 139 | |
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 140 | code_const "op < \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 141 | (Haskell infix 4 "<") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 142 | (Scala infixl 4 "<") | 
| 
9728342bcd56
another refinement chapter in the neverending numeral story
 haftmann parents: 
37947diff
changeset | 143 | |
| 37968 | 144 | end |