src/HOL/Library/Code_Natural.thy
author wenzelm
Sat, 07 Apr 2012 16:41:59 +0200
changeset 47389 e8552cba702d
parent 47108 2a1953f0d20d
permissions -rw-r--r--
explicit checks stable_finished_theory/stable_command allow parallel asynchronous command transactions; tuned;
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
46547
d1dcb91a512e use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents: 39781
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"
38810
361119ea62ee re-added accidental omission
haftmann
parents: 38781
diff changeset
    12
{*import Data.Array.ST;
361119ea62ee re-added accidental omission
haftmann
parents: 38781
diff changeset
    13
361119ea62ee re-added accidental omission
haftmann
parents: 38781
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
38810
361119ea62ee re-added accidental omission
haftmann
parents: 38781
diff changeset
    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
567b94f8bb6e tuned includes
haftmann
parents: 38771
diff changeset
    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: 38857
diff changeset
    59
{*object Natural {
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    60
38968
e55deaa22fff do not print object frame around Scala includes -- this is in the responsibility of the user
haftmann
parents: 38857
diff 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: 38857
diff 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: 38857
diff changeset
    63
  def apply(numeral: String): Natural = Natural(BigInt(numeral))
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    64
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
38968
e55deaa22fff do not print object frame around Scala includes -- this is in the responsibility of the user
haftmann
parents: 38857
diff changeset
    67
class Natural private(private val value: BigInt) {
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    68
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    69
  override def hashCode(): Int = this.value.hashCode()
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 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: 38857
diff changeset
    72
    case that: Natural => this equals that
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    73
    case _ => false
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    74
  }
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
  override def toString(): String = this.value.toString
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    77
38968
e55deaa22fff do not print object frame around Scala includes -- this is in the responsibility of the user
haftmann
parents: 38857
diff changeset
    78
  def equals(that: Natural): Boolean = this.value == that.value
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    79
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    80
  def as_BigInt: BigInt = this.value
39781
2053638a2bf2 scala is reserved identifier
haftmann
parents: 39272
diff changeset
    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: 37947
diff changeset
    82
      this.value.intValue
37968
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
    83
    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
    84
38968
e55deaa22fff do not print object frame around Scala includes -- this is in the responsibility of the user
haftmann
parents: 38857
diff 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: 38857
diff 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: 38857
diff changeset
    87
  def *(that: Natural): Natural = new Natural(this.value * that.value)
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    88
38968
e55deaa22fff do not print object frame around Scala includes -- this is in the responsibility of the user
haftmann
parents: 38857
diff 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: 37947
diff changeset
    90
    else {
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff 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: 38857
diff changeset
    92
      (new Natural(k), new Natural(l))
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    93
    }
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    94
38968
e55deaa22fff do not print object frame around Scala includes -- this is in the responsibility of the user
haftmann
parents: 38857
diff changeset
    95
  def <=(that: Natural): Boolean = this.value <= that.value
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    96
38968
e55deaa22fff do not print object frame around Scala includes -- this is in the responsibility of the user
haftmann
parents: 38857
diff changeset
    97
  def <(that: Natural): Boolean = this.value < that.value
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
    98
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
code_reserved Scala Natural
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_type code_numeral
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff 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: 38857
diff changeset
   106
  (Scala "Natural")
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   107
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   108
setup {*
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   109
  fold (Numeral.add_code @{const_name Code_Numeral.Num}
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   110
    false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"]
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   111
*}
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   112
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38810
diff changeset
   113
code_instance code_numeral :: equal
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   114
  (Haskell -)
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   115
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   116
code_const "0::code_numeral"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   117
  (Haskell "0")
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   118
  (Scala "Natural(0)")
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   119
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   120
code_const "plus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   121
  (Haskell infixl 6 "+")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   122
  (Scala infixl 7 "+")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   123
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   124
code_const "minus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   125
  (Haskell infixl 6 "-")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   126
  (Scala infixl 7 "-")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   127
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   128
code_const "times \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   129
  (Haskell infixl 7 "*")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   130
  (Scala infixl 8 "*")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   131
46547
d1dcb91a512e use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents: 39781
diff changeset
   132
code_const Code_Numeral.div_mod
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   133
  (Haskell "divMod")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   134
  (Scala infixl 8 "/%")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   135
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38810
diff changeset
   136
code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
39272
0b61951d2682 Haskell == is infix, not infixl
haftmann
parents: 38968
diff changeset
   137
  (Haskell infix 4 "==")
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   138
  (Scala infixl 5 "==")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   139
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   140
code_const "less_eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   141
  (Haskell infix 4 "<=")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   142
  (Scala infixl 4 "<=")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   143
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46547
diff changeset
   144
code_const "less \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
37958
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   145
  (Haskell infix 4 "<")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   146
  (Scala infixl 4 "<")
9728342bcd56 another refinement chapter in the neverending numeral story
haftmann
parents: 37947
diff changeset
   147
37968
52fdcb76c0af added Code_Natural.thy
haftmann
parents: 37958
diff changeset
   148
end