src/HOL/Library/Code_Natural.thy
author haftmann
Thu, 26 Aug 2010 13:25:14 +0200
changeset 38775 741ca0c98f6f
parent 38774 567b94f8bb6e
child 38781 6b356e3687d2
permissions -rw-r--r--
re-added accidental omission
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37968
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
     1
(*  Title:      HOL/Library/Code_Natural.thy
24999
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
haftmann
parents:
diff changeset
     3
*)
haftmann
parents:
diff changeset
     4
37968
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
     5
theory Code_Natural
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31192
diff changeset
     6
imports Main
24999
haftmann
parents:
diff changeset
     7
begin
haftmann
parents:
diff changeset
     8
37968
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
     9
section {* Alternative representation of @{typ code_numeral} for @{text Haskell} and @{text Scala} *}
24999
haftmann
parents:
diff changeset
    10
37968
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    11
code_include Haskell "Natural"
38775
741ca0c98f6f re-added accidental omission
haftmann
parents: 38774
diff changeset
    12
{*import Data.Array.ST;
741ca0c98f6f re-added accidental omission
haftmann
parents: 38774
diff changeset
    13
741ca0c98f6f re-added accidental omission
haftmann
parents: 38774
diff changeset
    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
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    31
instance Ix Natural where {
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    32
  range (Natural n, Natural m) = map Natural (range (n, m));
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    33
  index (Natural n, Natural m) (Natural q) = index (n, m) q;
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    34
  inRange (Natural n, Natural m) (Natural q) = inRange (n, m) q;
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    35
  rangeSize (Natural n, Natural m) = rangeSize (n, m);
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    36
};
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    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
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    53
};*}
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    54
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    55
code_reserved Haskell Natural
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    56
38774
567b94f8bb6e tuned includes
haftmann
parents: 38771
diff changeset
    57
code_include Scala "Natural"
567b94f8bb6e tuned includes
haftmann
parents: 38771
diff changeset
    58
{*import scala.Math
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    59
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    60
object Nat {
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    61
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    62
  def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    63
  def apply(numeral: Int): Nat = Nat(BigInt(numeral))
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    64
  def apply(numeral: String): Nat = Nat(BigInt(numeral))
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    65
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
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    68
class Nat private(private val value: BigInt) {
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    69
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    70
  override def hashCode(): Int = this.value.hashCode()
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    71
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    72
  override def equals(that: Any): Boolean = that match {
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    73
    case that: Nat => this equals that
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    74
    case _ => false
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    75
  }
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
  override def toString(): String = this.value.toString
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    78
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    79
  def equals(that: Nat): Boolean = this.value == that.value
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    80
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    81
  def as_BigInt: BigInt = this.value
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    82
  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
    83
      this.value.intValue
37968
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    84
    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
    85
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    86
  def +(that: Nat): Nat = new Nat(this.value + that.value)
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    87
  def -(that: Nat): Nat = Nat(this.value - that.value)
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    88
  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
    89
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    90
  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
    91
    else {
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    92
      val (k, l) = this.value /% that.value
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    93
      (new Nat(k), new Nat(l))
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    94
    }
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    95
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    96
  def <=(that: Nat): Boolean = this.value <= that.value
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    97
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
    98
  def <(that: Nat): Boolean = this.value < that.value
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    99
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
code_reserved Scala Natural
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   104
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   105
code_type code_numeral
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   106
  (Haskell "Natural.Natural")
38771
f9cd27cbe8a4 code_include Scala: qualify module nmae
haftmann
parents: 37968
diff changeset
   107
  (Scala "Natural.Nat")
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   108
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   109
setup {*
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   110
  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
   111
    false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"]
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   112
*}
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
code_instance code_numeral :: eq
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   115
  (Haskell -)
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   116
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   117
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
   118
  (Haskell infixl 6 "+")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   119
  (Scala infixl 7 "+")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   120
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   121
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
   122
  (Haskell infixl 6 "-")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   123
  (Scala infixl 7 "-")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   124
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   125
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
   126
  (Haskell infixl 7 "*")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   127
  (Scala infixl 8 "*")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   128
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   129
code_const div_mod_code_numeral
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   130
  (Haskell "divMod")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   131
  (Scala infixl 8 "/%")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   132
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   133
code_const "eq_class.eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   134
  (Haskell infixl 4 "==")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   135
  (Scala infixl 5 "==")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   136
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   137
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
   138
  (Haskell infix 4 "<=")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   139
  (Scala infixl 4 "<=")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   140
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   141
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
   142
  (Haskell infix 4 "<")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   143
  (Scala infixl 4 "<")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   144
37968
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
   145
end